src/HOL/Extraction/Higman.thy
author berghofe
Thu, 25 Aug 2005 16:13:09 +0200
changeset 17145 e623e57b0f44
parent 16417 9bc16273c2d4
child 17604 5f30179fbf44
permissions -rw-r--r--
Adapted to new code generator syntax.
berghofe@13405
     1
(*  Title:      HOL/Extraction/Higman.thy
berghofe@13405
     2
    ID:         $Id$
berghofe@13405
     3
    Author:     Stefan Berghofer, TU Muenchen
berghofe@13405
     4
                Monika Seisenberger, LMU Muenchen
berghofe@13405
     5
*)
berghofe@13405
     6
berghofe@13405
     7
header {* Higman's lemma *}
berghofe@13405
     8
haftmann@16417
     9
theory Higman imports Main begin
berghofe@13405
    10
berghofe@13405
    11
text {*
berghofe@13405
    12
  Formalization by Stefan Berghofer and Monika Seisenberger,
berghofe@13405
    13
  based on Coquand and Fridlender \cite{Coquand93}.
berghofe@13405
    14
*}
berghofe@13405
    15
berghofe@13405
    16
datatype letter = A | B
berghofe@13405
    17
berghofe@13405
    18
consts
berghofe@13405
    19
  emb :: "(letter list \<times> letter list) set"
berghofe@13405
    20
berghofe@13405
    21
inductive emb
berghofe@13405
    22
intros
wenzelm@15801
    23
  emb0 [Pure.intro]: "([], bs) \<in> emb"
wenzelm@15801
    24
  emb1 [Pure.intro]: "(as, bs) \<in> emb \<Longrightarrow> (as, b # bs) \<in> emb"
wenzelm@15801
    25
  emb2 [Pure.intro]: "(as, bs) \<in> emb \<Longrightarrow> (a # as, a # bs) \<in> emb"
berghofe@13405
    26
berghofe@13405
    27
consts
berghofe@13405
    28
  L :: "letter list \<Rightarrow> letter list list set"
berghofe@13405
    29
berghofe@13969
    30
inductive "L v"
berghofe@13405
    31
intros
wenzelm@15801
    32
  L0 [Pure.intro]: "(w, v) \<in> emb \<Longrightarrow> w # ws \<in> L v"
wenzelm@15801
    33
  L1 [Pure.intro]: "ws \<in> L v \<Longrightarrow> w # ws \<in> L v"
berghofe@13405
    34
berghofe@13405
    35
consts
berghofe@13405
    36
  good :: "letter list list set"
berghofe@13405
    37
berghofe@13405
    38
inductive good
berghofe@13405
    39
intros
wenzelm@15801
    40
  good0 [Pure.intro]: "ws \<in> L w \<Longrightarrow> w # ws \<in> good"
wenzelm@15801
    41
  good1 [Pure.intro]: "ws \<in> good \<Longrightarrow> w # ws \<in> good"
berghofe@13405
    42
berghofe@13405
    43
consts
berghofe@13405
    44
  R :: "letter \<Rightarrow> (letter list list \<times> letter list list) set"
berghofe@13405
    45
berghofe@13405
    46
inductive "R a"
berghofe@13405
    47
intros
wenzelm@15801
    48
  R0 [Pure.intro]: "([], []) \<in> R a"
wenzelm@15801
    49
  R1 [Pure.intro]: "(vs, ws) \<in> R a \<Longrightarrow> (w # vs, (a # w) # ws) \<in> R a"
berghofe@13405
    50
berghofe@13405
    51
consts
berghofe@13405
    52
  T :: "letter \<Rightarrow> (letter list list \<times> letter list list) set"
berghofe@13405
    53
berghofe@13405
    54
inductive "T a"
berghofe@13405
    55
intros
wenzelm@15801
    56
  T0 [Pure.intro]: "a \<noteq> b \<Longrightarrow> (ws, zs) \<in> R b \<Longrightarrow> (w # zs, (a # w) # zs) \<in> T a"
wenzelm@15801
    57
  T1 [Pure.intro]: "(ws, zs) \<in> T a \<Longrightarrow> (w # ws, (a # w) # zs) \<in> T a"
wenzelm@15801
    58
  T2 [Pure.intro]: "a \<noteq> b \<Longrightarrow> (ws, zs) \<in> T a \<Longrightarrow> (ws, (b # w) # zs) \<in> T a"
berghofe@13405
    59
berghofe@13405
    60
consts
berghofe@13405
    61
  bar :: "letter list list set"
berghofe@13405
    62
berghofe@13405
    63
inductive bar
berghofe@13405
    64
intros
wenzelm@15801
    65
  bar1 [Pure.intro]: "ws \<in> good \<Longrightarrow> ws \<in> bar"
wenzelm@15801
    66
  bar2 [Pure.intro]: "(\<And>w. w # ws \<in> bar) \<Longrightarrow> ws \<in> bar"
berghofe@13405
    67
berghofe@13405
    68
theorem prop1: "([] # ws) \<in> bar" by rules
berghofe@13405
    69
berghofe@13405
    70
theorem lemma1: "ws \<in> L as \<Longrightarrow> ws \<in> L (a # as)"
berghofe@13405
    71
  by (erule L.induct, rules+)
berghofe@13405
    72
berghofe@13969
    73
lemma lemma2': "(vs, ws) \<in> R a \<Longrightarrow> vs \<in> L as \<Longrightarrow> ws \<in> L (a # as)"
berghofe@13969
    74
  apply (induct set: R)
berghofe@13405
    75
  apply (erule L.elims)
berghofe@13405
    76
  apply simp+
berghofe@13405
    77
  apply (erule L.elims)
berghofe@13405
    78
  apply simp_all
berghofe@13405
    79
  apply (rule L0)
berghofe@13405
    80
  apply (erule emb2)
berghofe@13405
    81
  apply (erule L1)
berghofe@13405
    82
  done
berghofe@13969
    83
berghofe@13969
    84
lemma lemma2: "(vs, ws) \<in> R a \<Longrightarrow> vs \<in> good \<Longrightarrow> ws \<in> good"
berghofe@13969
    85
  apply (induct set: R)
berghofe@13405
    86
  apply rules
berghofe@13405
    87
  apply (erule good.elims)
berghofe@13405
    88
  apply simp_all
berghofe@13405
    89
  apply (rule good0)
berghofe@13405
    90
  apply (erule lemma2')
berghofe@13405
    91
  apply assumption
berghofe@13405
    92
  apply (erule good1)
berghofe@13405
    93
  done
berghofe@13405
    94
berghofe@13969
    95
lemma lemma3': "(vs, ws) \<in> T a \<Longrightarrow> vs \<in> L as \<Longrightarrow> ws \<in> L (a # as)"
berghofe@13969
    96
  apply (induct set: T)
berghofe@13405
    97
  apply (erule L.elims)
berghofe@13405
    98
  apply simp_all
berghofe@13405
    99
  apply (rule L0)
berghofe@13405
   100
  apply (erule emb2)
berghofe@13405
   101
  apply (rule L1)
berghofe@13405
   102
  apply (erule lemma1)
berghofe@13405
   103
  apply (erule L.elims)
berghofe@13405
   104
  apply simp_all
berghofe@13405
   105
  apply rules+
berghofe@13405
   106
  done
berghofe@13405
   107
berghofe@13969
   108
lemma lemma3: "(ws, zs) \<in> T a \<Longrightarrow> ws \<in> good \<Longrightarrow> zs \<in> good"
berghofe@13969
   109
  apply (induct set: T)
berghofe@13405
   110
  apply (erule good.elims)
berghofe@13405
   111
  apply simp_all
berghofe@13405
   112
  apply (rule good0)
berghofe@13405
   113
  apply (erule lemma1)
berghofe@13405
   114
  apply (erule good1)
berghofe@13405
   115
  apply (erule good.elims)
berghofe@13405
   116
  apply simp_all
berghofe@13405
   117
  apply (rule good0)
berghofe@13405
   118
  apply (erule lemma3')
berghofe@13405
   119
  apply rules+
berghofe@13405
   120
  done
berghofe@13405
   121
berghofe@13969
   122
lemma lemma4: "(ws, zs) \<in> R a \<Longrightarrow> ws \<noteq> [] \<Longrightarrow> (ws, zs) \<in> T a"
berghofe@13969
   123
  apply (induct set: R)
berghofe@13405
   124
  apply rules
berghofe@13405
   125
  apply (case_tac vs)
berghofe@13405
   126
  apply (erule R.elims)
berghofe@13405
   127
  apply simp
berghofe@13405
   128
  apply (case_tac a)
berghofe@13405
   129
  apply (rule_tac b=B in T0)
berghofe@13405
   130
  apply simp
berghofe@13405
   131
  apply (rule R0)
berghofe@13405
   132
  apply (rule_tac b=A in T0)
berghofe@13405
   133
  apply simp
berghofe@13405
   134
  apply (rule R0)
berghofe@13405
   135
  apply simp
berghofe@13405
   136
  apply (rule T1)
berghofe@13405
   137
  apply simp
berghofe@13405
   138
  done
berghofe@13405
   139
berghofe@13930
   140
lemma letter_neq: "(a::letter) \<noteq> b \<Longrightarrow> c \<noteq> a \<Longrightarrow> c = b"
berghofe@13930
   141
  apply (case_tac a)
berghofe@13930
   142
  apply (case_tac b)
berghofe@13930
   143
  apply (case_tac c, simp, simp)
berghofe@13930
   144
  apply (case_tac c, simp, simp)
berghofe@13930
   145
  apply (case_tac b)
berghofe@13930
   146
  apply (case_tac c, simp, simp)
berghofe@13930
   147
  apply (case_tac c, simp, simp)
berghofe@13930
   148
  done
berghofe@13405
   149
berghofe@13930
   150
lemma letter_eq_dec: "(a::letter) = b \<or> a \<noteq> b"
berghofe@13405
   151
  apply (case_tac a)
berghofe@13405
   152
  apply (case_tac b)
berghofe@13405
   153
  apply simp
berghofe@13405
   154
  apply simp
berghofe@13405
   155
  apply (case_tac b)
berghofe@13405
   156
  apply simp
berghofe@13405
   157
  apply simp
berghofe@13405
   158
  done
berghofe@13405
   159
berghofe@13930
   160
theorem prop2:
berghofe@13930
   161
  assumes ab: "a \<noteq> b" and bar: "xs \<in> bar"
berghofe@13930
   162
  shows "\<And>ys zs. ys \<in> bar \<Longrightarrow> (xs, zs) \<in> T a \<Longrightarrow> (ys, zs) \<in> T b \<Longrightarrow> zs \<in> bar" using bar
berghofe@13930
   163
proof induct
berghofe@13930
   164
  fix xs zs assume "xs \<in> good" and "(xs, zs) \<in> T a"
berghofe@13930
   165
  show "zs \<in> bar" by (rule bar1) (rule lemma3)
berghofe@13930
   166
next
berghofe@13930
   167
  fix xs ys
berghofe@13930
   168
  assume I: "\<And>w ys zs. ys \<in> bar \<Longrightarrow> (w # xs, zs) \<in> T a \<Longrightarrow> (ys, zs) \<in> T b \<Longrightarrow> zs \<in> bar"
berghofe@13930
   169
  assume "ys \<in> bar"
berghofe@13930
   170
  thus "\<And>zs. (xs, zs) \<in> T a \<Longrightarrow> (ys, zs) \<in> T b \<Longrightarrow> zs \<in> bar"
berghofe@13930
   171
  proof induct
berghofe@13930
   172
    fix ys zs assume "ys \<in> good" and "(ys, zs) \<in> T b"
berghofe@13930
   173
    show "zs \<in> bar" by (rule bar1) (rule lemma3)
berghofe@13930
   174
  next
berghofe@13930
   175
    fix ys zs assume I': "\<And>w zs. (xs, zs) \<in> T a \<Longrightarrow> (w # ys, zs) \<in> T b \<Longrightarrow> zs \<in> bar"
berghofe@13930
   176
    and ys: "\<And>w. w # ys \<in> bar" and Ta: "(xs, zs) \<in> T a" and Tb: "(ys, zs) \<in> T b"
berghofe@13930
   177
    show "zs \<in> bar"
berghofe@13930
   178
    proof (rule bar2)
berghofe@13930
   179
      fix w
berghofe@13930
   180
      show "w # zs \<in> bar"
berghofe@13930
   181
      proof (cases w)
berghofe@13930
   182
	case Nil
berghofe@13930
   183
	thus ?thesis by simp (rule prop1)
berghofe@13930
   184
      next
berghofe@13930
   185
	case (Cons c cs)
berghofe@13930
   186
	from letter_eq_dec show ?thesis
berghofe@13930
   187
	proof
berghofe@13930
   188
	  assume ca: "c = a"
berghofe@13969
   189
	  from ab have "(a # cs) # zs \<in> bar" by (rules intro: I ys Ta Tb)
berghofe@13930
   190
	  thus ?thesis by (simp add: Cons ca)
berghofe@13930
   191
	next
berghofe@13930
   192
	  assume "c \<noteq> a"
berghofe@13930
   193
	  with ab have cb: "c = b" by (rule letter_neq)
berghofe@13969
   194
	  from ab have "(b # cs) # zs \<in> bar" by (rules intro: I' Ta Tb)
berghofe@13930
   195
	  thus ?thesis by (simp add: Cons cb)
berghofe@13930
   196
	qed
berghofe@13930
   197
      qed
berghofe@13930
   198
    qed
berghofe@13930
   199
  qed
berghofe@13930
   200
qed
berghofe@13405
   201
berghofe@13930
   202
theorem prop3:
berghofe@13930
   203
  assumes bar: "xs \<in> bar"
berghofe@13930
   204
  shows "\<And>zs. xs \<noteq> [] \<Longrightarrow> (xs, zs) \<in> R a \<Longrightarrow> zs \<in> bar" using bar
berghofe@13930
   205
proof induct
berghofe@13930
   206
  fix xs zs
berghofe@13930
   207
  assume "xs \<in> good" and "(xs, zs) \<in> R a"
berghofe@13930
   208
  show "zs \<in> bar" by (rule bar1) (rule lemma2)
berghofe@13930
   209
next
berghofe@13930
   210
  fix xs zs
berghofe@13930
   211
  assume I: "\<And>w zs. w # xs \<noteq> [] \<Longrightarrow> (w # xs, zs) \<in> R a \<Longrightarrow> zs \<in> bar"
berghofe@13930
   212
  and xsb: "\<And>w. w # xs \<in> bar" and xsn: "xs \<noteq> []" and R: "(xs, zs) \<in> R a"
berghofe@13930
   213
  show "zs \<in> bar"
berghofe@13930
   214
  proof (rule bar2)
berghofe@13930
   215
    fix w
berghofe@13930
   216
    show "w # zs \<in> bar"
berghofe@13930
   217
    proof (induct w)
berghofe@13930
   218
      case Nil
berghofe@13930
   219
      show ?case by (rule prop1)
berghofe@13930
   220
    next
berghofe@13930
   221
      case (Cons c cs)
berghofe@13930
   222
      from letter_eq_dec show ?case
berghofe@13930
   223
      proof
berghofe@13930
   224
	assume "c = a"
berghofe@13930
   225
	thus ?thesis by (rules intro: I [simplified] R)
berghofe@13930
   226
      next
berghofe@13930
   227
	from R xsn have T: "(xs, zs) \<in> T a" by (rule lemma4)
berghofe@13930
   228
	assume "c \<noteq> a"
berghofe@13930
   229
	thus ?thesis by (rules intro: prop2 Cons xsb xsn R T)
berghofe@13930
   230
      qed
berghofe@13930
   231
    qed
berghofe@13930
   232
  qed
berghofe@13930
   233
qed
berghofe@13405
   234
berghofe@13405
   235
theorem higman: "[] \<in> bar"
berghofe@13930
   236
proof (rule bar2)
berghofe@13930
   237
  fix w
berghofe@13930
   238
  show "[w] \<in> bar"
berghofe@13930
   239
  proof (induct w)
berghofe@13930
   240
    show "[[]] \<in> bar" by (rule prop1)
berghofe@13930
   241
  next
berghofe@13930
   242
    fix c cs assume "[cs] \<in> bar"
berghofe@13930
   243
    thus "[c # cs] \<in> bar" by (rule prop3) (simp, rules)
berghofe@13930
   244
  qed
berghofe@13930
   245
qed
berghofe@13405
   246
berghofe@13405
   247
consts
berghofe@13405
   248
  is_prefix :: "'a list \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> bool"
berghofe@13405
   249
berghofe@13405
   250
primrec
berghofe@13405
   251
  "is_prefix [] f = True"
berghofe@13405
   252
  "is_prefix (x # xs) f = (x = f (length xs) \<and> is_prefix xs f)"
berghofe@13405
   253
berghofe@13405
   254
theorem good_prefix_lemma:
berghofe@13930
   255
  assumes bar: "ws \<in> bar"
berghofe@13930
   256
  shows "is_prefix ws f \<Longrightarrow> \<exists>vs. is_prefix vs f \<and> vs \<in> good" using bar
berghofe@13930
   257
proof induct
berghofe@13930
   258
  case bar1
berghofe@13930
   259
  thus ?case by rules
berghofe@13930
   260
next
berghofe@13930
   261
  case (bar2 ws)
berghofe@13930
   262
  have "is_prefix (f (length ws) # ws) f" by simp
berghofe@13930
   263
  thus ?case by (rules intro: bar2)
berghofe@13930
   264
qed
berghofe@13405
   265
berghofe@13405
   266
theorem good_prefix: "\<exists>vs. is_prefix vs f \<and> vs \<in> good"
berghofe@13930
   267
  using higman
berghofe@13930
   268
  by (rule good_prefix_lemma) simp+
berghofe@13405
   269
berghofe@13711
   270
subsection {* Extracting the program *}
berghofe@13405
   271
berghofe@13711
   272
declare bar.induct [ind_realizer]
berghofe@13405
   273
berghofe@13405
   274
extract good_prefix
berghofe@13405
   275
berghofe@13405
   276
text {*
berghofe@13405
   277
  Program extracted from the proof of @{text good_prefix}:
berghofe@13405
   278
  @{thm [display] good_prefix_def [no_vars]}
berghofe@13405
   279
  Corresponding correctness theorem:
berghofe@13405
   280
  @{thm [display] good_prefix_correctness [no_vars]}
berghofe@13405
   281
  Program extracted from the proof of @{text good_prefix_lemma}:
berghofe@13405
   282
  @{thm [display] good_prefix_lemma_def [no_vars]}
berghofe@13405
   283
  Program extracted from the proof of @{text higman}:
berghofe@13405
   284
  @{thm [display] higman_def [no_vars]}
berghofe@13405
   285
  Program extracted from the proof of @{text prop1}:
berghofe@13405
   286
  @{thm [display] prop1_def [no_vars]}
berghofe@13405
   287
  Program extracted from the proof of @{text prop2}:
berghofe@13405
   288
  @{thm [display] prop2_def [no_vars]}
berghofe@13405
   289
  Program extracted from the proof of @{text prop3}:
berghofe@13405
   290
  @{thm [display] prop3_def [no_vars]}
berghofe@13405
   291
*}
berghofe@13405
   292
berghofe@17145
   293
code_module Higman
berghofe@17145
   294
contains
berghofe@13405
   295
  test = good_prefix
berghofe@13405
   296
berghofe@13405
   297
ML {*
berghofe@17145
   298
local open Higman in
berghofe@17145
   299
berghofe@13405
   300
val a = 16807.0;
berghofe@13405
   301
val m = 2147483647.0;
berghofe@13405
   302
berghofe@13405
   303
fun nextRand seed =
berghofe@13405
   304
  let val t = a*seed
berghofe@13405
   305
  in  t - m * real (Real.floor(t/m)) end;
berghofe@13405
   306
berghofe@13405
   307
fun mk_word seed l =
berghofe@13405
   308
  let
berghofe@13405
   309
    val r = nextRand seed;
berghofe@13405
   310
    val i = Real.round (r / m * 10.0);
berghofe@13405
   311
  in if i > 7 andalso l > 2 then (r, []) else
berghofe@13405
   312
    apsnd (cons (if i mod 2 = 0 then A else B)) (mk_word r (l+1))
berghofe@13405
   313
  end;
berghofe@13405
   314
berghofe@17145
   315
fun f s id_0 = mk_word s 0
berghofe@13405
   316
  | f s (Suc n) = f (fst (mk_word s 0)) n;
berghofe@13405
   317
berghofe@13405
   318
val g1 = snd o (f 20000.0);
berghofe@13405
   319
berghofe@13405
   320
val g2 = snd o (f 50000.0);
berghofe@13405
   321
berghofe@17145
   322
fun f1 id_0 = [A,A]
berghofe@17145
   323
  | f1 (Suc id_0) = [B]
berghofe@17145
   324
  | f1 (Suc (Suc id_0)) = [A,B]
berghofe@13405
   325
  | f1 _ = [];
berghofe@13405
   326
berghofe@17145
   327
fun f2 id_0 = [A,A]
berghofe@17145
   328
  | f2 (Suc id_0) = [B]
berghofe@17145
   329
  | f2 (Suc (Suc id_0)) = [B,A]
berghofe@13405
   330
  | f2 _ = [];
berghofe@13405
   331
berghofe@13405
   332
val xs1 = test g1;
berghofe@13405
   333
val xs2 = test g2;
berghofe@13405
   334
val xs3 = test f1;
berghofe@13405
   335
val xs4 = test f2;
berghofe@17145
   336
berghofe@17145
   337
end;
berghofe@13405
   338
*}
berghofe@13405
   339
berghofe@13405
   340
end