Closures2.thy
author urbanc
Thu, 01 Sep 2011 23:18:34 +0000
changeset 222 191769fc68c3
parent 221 68e28debe995
child 223 5f7b7ad498dd
permissions -rw-r--r--
included Higman's lemma from the Isabelle repository
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
     1
theory Closure2
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
222
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
     4
  Higman
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
     5
  (* "~~/src/HOL/Proofs/Extraction/Higman" *)
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
     6
begin
91e3e906034c added a further test
urbanc
parents:
diff changeset
     7
222
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
     8
notation
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
     9
  emb ("_ \<preceq> _")
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    10
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    11
declare  emb0 [intro]
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    12
declare  emb1 [intro]
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    13
declare  emb2 [intro]
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    14
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    15
lemma letter_UNIV:
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    16
  shows "UNIV = {A, B::letter}"
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    17
apply(auto)
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    18
apply(case_tac x)
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    19
apply(auto)
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    20
done
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    21
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    22
instance letter :: finite
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    23
apply(default)
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    24
apply(simp add: letter_UNIV)
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    25
done
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    26
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    27
hide_const A
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    28
hide_const B
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    29
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    30
(*
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    31
inductive 
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    32
  emb :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" ("_ \<preceq> _")
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
    33
where
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    34
   emb0 [intro]: "emb [] y"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    35
 | emb1 [intro]: "emb x y \<Longrightarrow> emb x (c # y)"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    36
 | 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
    37
*)
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
    38
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    39
lemma emb_refl [intro]:
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
    40
  shows "x \<preceq> x"
222
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    41
by (induct x) (auto)
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
    42
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    43
lemma emb_right [intro]:
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
    44
  assumes a: "x \<preceq> y"
91e3e906034c added a further test
urbanc
parents:
diff changeset
    45
  shows "x \<preceq> y @ y'"
222
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    46
using a 
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
    47
by (induct arbitrary: y') (auto)
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
    48
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    49
lemma emb_left [intro]:
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
    50
  assumes a: "x \<preceq> y"
91e3e906034c added a further test
urbanc
parents:
diff changeset
    51
  shows "x \<preceq> y' @ y"
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    52
using a by (induct 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_appendI [intro]:
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
    55
  assumes a: "x \<preceq> x'"
91e3e906034c added a further test
urbanc
parents:
diff changeset
    56
  and     b: "y \<preceq> y'"
91e3e906034c added a further test
urbanc
parents:
diff changeset
    57
  shows "x @ y \<preceq> x' @ y'"
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    58
using a b by (induct) (auto)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    59
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    60
lemma emb_cons_leftD:
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    61
  assumes "a # x \<preceq> y"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    62
  shows "\<exists>y1 y2. y = y1 @ [a] @ y2 \<and> x \<preceq> y2"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    63
using assms
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    64
apply(induct x\<equiv>"a # x" y\<equiv>"y" arbitrary: a x y)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    65
apply(auto)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    66
apply(metis append_Cons)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    67
done
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    68
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    69
lemma emb_append_leftD:
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    70
  assumes "x1 @ x2 \<preceq> y"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    71
  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
    72
using assms
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    73
apply(induct x1 arbitrary: x2 y)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    74
apply(auto)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    75
apply(drule emb_cons_leftD)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    76
apply(auto)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    77
apply(drule_tac x="x2" in meta_spec)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    78
apply(drule_tac x="y2" in meta_spec)
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(rule_tac x="y1 @ (a # y1a)" in exI)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    81
apply(rule_tac x="y2a" in exI)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    82
apply(auto)
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
    83
done
91e3e906034c added a further test
urbanc
parents:
diff changeset
    84
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    85
lemma emb_trans:
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    86
  assumes a: "x1 \<preceq> x2"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    87
  and     b: "x2 \<preceq> x3"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    88
  shows "x1 \<preceq> x3"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    89
using a b
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    90
apply(induct arbitrary: x3)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    91
apply(metis emb0)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    92
apply(metis emb_cons_leftD emb_left)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    93
apply(drule_tac emb_cons_leftD)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    94
apply(auto)
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
    95
done
91e3e906034c added a further test
urbanc
parents:
diff changeset
    96
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    97
lemma emb_strict_length:
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    98
  assumes a: "x \<preceq> y" "x \<noteq> y" 
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
    99
  shows "length x < length y"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   100
using a
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   101
by (induct) (auto simp add: less_Suc_eq)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   102
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   103
lemma emb_antisym:
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   104
  assumes a: "x \<preceq> y" "y \<preceq> x" 
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   105
  shows "x = y"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   106
using a emb_strict_length
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   107
by (metis not_less_iff_gr_or_eq)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   108
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   109
lemma emb_wf:
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   110
  shows "wf {(x, y). x \<preceq> y \<and> x \<noteq> y}"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   111
proof -
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   112
  have "wf (measure length)" by simp
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   113
  moreover
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   114
  have "{(x, y). x \<preceq> y \<and> x \<noteq> y} \<subseteq> measure length"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   115
    unfolding measure_def by (auto simp add: emb_strict_length)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   116
  ultimately 
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   117
  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
   118
qed
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   119
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   120
subsection {* Sub- and Supsequences *}
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
   121
91e3e906034c added a further test
urbanc
parents:
diff changeset
   122
definition
91e3e906034c added a further test
urbanc
parents:
diff changeset
   123
 "SUBSEQ A \<equiv> {x. \<exists>y \<in> A. x \<preceq> y}"
91e3e906034c added a further test
urbanc
parents:
diff changeset
   124
91e3e906034c added a further test
urbanc
parents:
diff changeset
   125
definition
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   126
 "SUPSEQ A \<equiv> {x. \<exists>y \<in> A. y \<preceq> x}"
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
   127
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   128
lemma SUPSEQ_simps [simp]:
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   129
  shows "SUPSEQ {} = {}"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   130
  and   "SUPSEQ {[]} = UNIV"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   131
unfolding SUPSEQ_def by auto
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
   132
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   133
lemma SUPSEQ_atom [simp]:
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   134
  shows "SUPSEQ {[a]} = UNIV \<cdot> {[a]} \<cdot> UNIV"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   135
unfolding SUPSEQ_def conc_def
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   136
by (auto) (metis append_Cons emb_cons_leftD)
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
   137
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   138
lemma SUPSEQ_union [simp]:
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   139
  shows "SUPSEQ (A \<union> B) = SUPSEQ A \<union> SUPSEQ B"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   140
unfolding SUPSEQ_def by auto
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
   141
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   142
lemma SUPSEQ_conc [simp]:
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   143
  shows "SUPSEQ (A \<cdot> B) = SUPSEQ A \<cdot> SUPSEQ B"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   144
unfolding SUPSEQ_def conc_def
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   145
by (auto) (metis emb_append_leftD)
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
   146
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   147
lemma SUPSEQ_star [simp]:
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   148
  shows "SUPSEQ (A\<star>) = UNIV"
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
   149
apply(subst star_unfold_left)
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   150
apply(simp only: SUPSEQ_union) 
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
   151
apply(simp)
91e3e906034c added a further test
urbanc
parents:
diff changeset
   152
done
91e3e906034c added a further test
urbanc
parents:
diff changeset
   153
91e3e906034c added a further test
urbanc
parents:
diff changeset
   154
definition
91e3e906034c added a further test
urbanc
parents:
diff changeset
   155
  Allreg :: "'a::finite rexp"
91e3e906034c added a further test
urbanc
parents:
diff changeset
   156
where
91e3e906034c added a further test
urbanc
parents:
diff changeset
   157
  "Allreg \<equiv> \<Uplus>(Atom ` UNIV)"
91e3e906034c added a further test
urbanc
parents:
diff changeset
   158
91e3e906034c added a further test
urbanc
parents:
diff changeset
   159
lemma Allreg_lang [simp]:
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   160
  shows "lang Allreg = (\<Union>a. {[a]})"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   161
unfolding Allreg_def by auto
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
   162
91e3e906034c added a further test
urbanc
parents:
diff changeset
   163
lemma [simp]:
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   164
  shows "(\<Union>a. {[a]})\<star> = UNIV"
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
   165
apply(auto)
91e3e906034c added a further test
urbanc
parents:
diff changeset
   166
apply(induct_tac x rule: list.induct)
91e3e906034c added a further test
urbanc
parents:
diff changeset
   167
apply(auto)
91e3e906034c added a further test
urbanc
parents:
diff changeset
   168
apply(subgoal_tac "[a] @ list \<in> (\<Union>a. {[a]})\<star>")
91e3e906034c added a further test
urbanc
parents:
diff changeset
   169
apply(simp)
91e3e906034c added a further test
urbanc
parents:
diff changeset
   170
apply(rule append_in_starI)
91e3e906034c added a further test
urbanc
parents:
diff changeset
   171
apply(auto)
91e3e906034c added a further test
urbanc
parents:
diff changeset
   172
done
91e3e906034c added a further test
urbanc
parents:
diff changeset
   173
91e3e906034c added a further test
urbanc
parents:
diff changeset
   174
lemma Star_Allreg_lang [simp]:
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   175
  shows "lang (Star Allreg) = UNIV"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   176
by simp
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
   177
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   178
fun 
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   179
  UP :: "'a::finite rexp \<Rightarrow> 'a rexp"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   180
where
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   181
  "UP (Zero) = Zero"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   182
| "UP (One) = Star Allreg"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   183
| "UP (Atom c) = Times (Star Allreg) (Times (Atom c) (Star Allreg))"   
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   184
| "UP (Plus r1 r2) = Plus (UP r1) (UP r2)"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   185
| "UP (Times r1 r2) = Times (UP r1) (UP r2)"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   186
| "UP (Star r) = Star Allreg"
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
   187
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   188
lemma lang_UP:
222
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   189
  fixes r::"letter rexp"
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   190
  shows "lang (UP r) = SUPSEQ (lang r)"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   191
by (induct r) (simp_all)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   192
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   193
lemma regular_SUPSEQ: 
222
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   194
  fixes A::"letter lang"
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
   195
  assumes "regular A"
91e3e906034c added a further test
urbanc
parents:
diff changeset
   196
  shows "regular (SUPSEQ A)"
91e3e906034c added a further test
urbanc
parents:
diff changeset
   197
proof -
222
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   198
  from assms obtain r::"letter rexp" where "lang r = A" by auto
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   199
  then have "lang (UP r) = SUPSEQ A" by (simp add: lang_UP)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   200
  then show "regular (SUPSEQ A)" by auto
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
   201
qed
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   202
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   203
lemma SUPSEQ_subset:
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   204
  shows "A \<subseteq> SUPSEQ A"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   205
unfolding SUPSEQ_def by auto
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   206
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   207
lemma w3:
222
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   208
  fixes T A::"letter lang"
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   209
  assumes eq: "T = - (SUBSEQ A)"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   210
  shows "T = SUPSEQ T"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   211
apply(rule)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   212
apply(rule SUPSEQ_subset)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   213
apply(rule ccontr)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   214
apply(auto)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   215
apply(rule ccontr)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   216
apply(subgoal_tac "x \<in> SUBSEQ A")
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   217
prefer 2
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   218
apply(subst (asm) (2) eq)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   219
apply(simp)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   220
apply(simp add: SUPSEQ_def)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   221
apply(erule bexE)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   222
apply(subgoal_tac "y \<in> SUBSEQ A")
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   223
prefer 2
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   224
apply(simp add: SUBSEQ_def)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   225
apply(erule bexE)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   226
apply(rule_tac x="xa" in bexI)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   227
apply(rule emb_trans)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   228
apply(assumption)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   229
apply(assumption)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   230
apply(assumption)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   231
apply(subst (asm) (2) eq)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   232
apply(simp)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   233
done
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   234
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   235
lemma w4:
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   236
  shows "- (SUBSEQ A) = SUPSEQ (- (SUBSEQ A))"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   237
by (rule w3) (simp)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   238
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   239
definition
222
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   240
  minimal_in :: "letter list \<Rightarrow> letter lang \<Rightarrow> bool"
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   241
where
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   242
  "minimal_in x A \<equiv> \<forall>y \<in> A. y \<preceq> x \<longrightarrow> y = x"
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   243
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   244
lemma minimal_in2:
222
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   245
  shows "minimal_in x A = (\<forall>y \<in> A. y \<preceq> x \<longrightarrow> x \<preceq> y)"
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   246
by (auto simp add: minimal_in_def intro: emb_antisym)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   247
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   248
lemma higman:
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   249
  assumes "\<forall>x \<in> A. \<forall>y \<in> A. x \<noteq> y \<longrightarrow> \<not>(x \<preceq> y) \<and> \<not>(y \<preceq> x)"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   250
  shows "finite A"
222
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   251
apply(rule ccontr)
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   252
apply(simp add: infinite_iff_countable_subset)
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   253
apply(auto)
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   254
apply(insert higman_idx)
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   255
apply(drule_tac x="f" in meta_spec)
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   256
apply(auto)
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   257
using assms
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   258
apply -
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   259
apply(drule_tac x="f i" in bspec)
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   260
apply(auto)[1]
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   261
apply(drule_tac x="f j" in bspec)
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   262
apply(auto)[1]
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   263
apply(drule mp)
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   264
apply(simp add: inj_on_def)
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   265
apply(auto)[1]
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   266
apply(auto)[1]
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   267
done
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   268
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   269
lemma minimal:
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   270
  assumes "minimal_in x A" "minimal_in y A"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   271
  and "x \<noteq> y" "x \<in> A" "y \<in> A"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   272
  shows "\<not>(x \<preceq> y) \<and> \<not>(y \<preceq> x)"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   273
using assms unfolding minimal_in_def 
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   274
by auto
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
   275
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   276
lemma main_lemma:
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   277
  "\<exists>M. M \<subseteq> A \<and> finite M \<and> SUPSEQ A = SUPSEQ M"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   278
proof -
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   279
  def M \<equiv> "{x \<in> A. minimal_in x A}"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   280
  have "M \<subseteq> A" unfolding M_def minimal_in_def by auto
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   281
  moreover
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   282
  have "finite M"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   283
    apply(rule higman)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   284
    unfolding M_def minimal_in_def
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   285
    by auto
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   286
  moreover
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   287
  have "SUPSEQ A \<subseteq> SUPSEQ M"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   288
    apply(rule)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   289
    apply(simp only: SUPSEQ_def)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   290
    apply(auto)[1]
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   291
    using emb_wf
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   292
    apply(erule_tac Q="{y' \<in> A. y' \<preceq> x}" and x="y" in wfE_min)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   293
    apply(simp)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   294
    apply(simp)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   295
    apply(rule_tac x="z" in bexI)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   296
    apply(simp)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   297
    apply(simp add: M_def)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   298
    apply(simp add: minimal_in2)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   299
    apply(rule ballI)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   300
    apply(rule impI)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   301
    apply(drule_tac x="ya" in meta_spec)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   302
    apply(simp)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   303
    apply(case_tac "ya = z")
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   304
    apply(auto)[1]
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   305
    apply(simp)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   306
    by (metis emb_trans)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   307
  moreover
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   308
  have "SUPSEQ M \<subseteq> SUPSEQ A"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   309
    by (auto simp add: SUPSEQ_def M_def)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   310
  ultimately
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   311
  show "\<exists>M. M \<subseteq> A \<and> finite M \<and> SUPSEQ A = SUPSEQ M" by blast
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   312
qed
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   313
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   314
lemma closure_SUPSEQ:
222
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   315
  fixes A::"letter lang" 
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   316
  shows "regular (SUPSEQ A)"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   317
proof -
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   318
  obtain M where a: "finite M" and b: "SUPSEQ A = SUPSEQ M"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   319
    using main_lemma by blast
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   320
  have "regular M" using a by (rule finite_regular)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   321
  then have "regular (SUPSEQ M)" by (rule regular_SUPSEQ)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   322
  then show "regular (SUPSEQ A)" using b by simp
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   323
qed
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   324
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   325
lemma closure_SUBSEQ:
222
191769fc68c3 included Higman's lemma from the Isabelle repository
urbanc
parents: 221
diff changeset
   326
  fixes A::"letter lang"
221
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   327
  shows "regular (SUBSEQ A)"
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   328
proof -
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   329
  have "regular (SUPSEQ (- SUBSEQ A))" by (rule closure_SUPSEQ)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   330
  then have "regular (- SUBSEQ A)" 
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   331
    apply(subst w4) 
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   332
    apply(assumption) 
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   333
    done
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   334
  then have "regular (- (- (SUBSEQ A)))" by (rule closure_complement)
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   335
  then show "regular (SUBSEQ A)" by simp
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   336
qed
68e28debe995 solved the SUBSEQ/SUPSEQ problem
urbanc
parents: 220
diff changeset
   337
220
91e3e906034c added a further test
urbanc
parents:
diff changeset
   338
91e3e906034c added a further test
urbanc
parents:
diff changeset
   339
91e3e906034c added a further test
urbanc
parents:
diff changeset
   340
end