src/HOL/Extraction/Higman.thy
author berghofe
Sun, 21 Jul 2002 15:44:42 +0200
changeset 13405 d20a4e67afc8
child 13470 d2cbbad84ad3
permissions -rw-r--r--
Examples for program extraction in HOL.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/Extraction/Higman.thy
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     2
    ID:         $Id$
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     3
    Author:     Stefan Berghofer, TU Muenchen
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     4
                Monika Seisenberger, LMU Muenchen
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     5
*)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     6
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     7
header {* Higman's lemma *}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     8
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     9
theory Higman = Main:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    10
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    11
text {*
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    12
  Formalization by Stefan Berghofer and Monika Seisenberger,
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    13
  based on Coquand and Fridlender \cite{Coquand93}.
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    14
*}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    15
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    16
datatype letter = A | B
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    17
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    18
consts
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    19
  emb :: "(letter list \<times> letter list) set"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    20
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    21
inductive emb
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    22
intros
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    23
  emb0 [CPure.intro]: "([], bs) \<in> emb"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    24
  emb1 [CPure.intro]: "(as, bs) \<in> emb \<Longrightarrow> (as, b # bs) \<in> emb"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    25
  emb2 [CPure.intro]: "(as, bs) \<in> emb \<Longrightarrow> (a # as, a # bs) \<in> emb"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    26
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    27
consts
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    28
  L :: "letter list \<Rightarrow> letter list list set"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    29
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    30
inductive "L y"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    31
intros
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    32
  L0 [CPure.intro]: "(w, y) \<in> emb \<Longrightarrow> w # ws \<in> L y"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    33
  L1 [CPure.intro]: "ws \<in> L y \<Longrightarrow> w # ws \<in> L y"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    34
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    35
consts
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    36
  good :: "letter list list set"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    37
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    38
inductive good
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    39
intros
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    40
  good0 [CPure.intro]: "ws \<in> L w \<Longrightarrow> w # ws \<in> good"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    41
  good1 [CPure.intro]: "ws \<in> good \<Longrightarrow> w # ws \<in> good"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    42
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    43
consts
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    44
  R :: "letter \<Rightarrow> (letter list list \<times> letter list list) set"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    45
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    46
inductive "R a"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    47
intros
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    48
  R0 [CPure.intro]: "([], []) \<in> R a"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    49
  R1 [CPure.intro]: "(vs, ws) \<in> R a \<Longrightarrow> (w # vs, (a # w) # ws) \<in> R a"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    50
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    51
consts
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    52
  T :: "letter \<Rightarrow> (letter list list \<times> letter list list) set"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    53
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    54
inductive "T a"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    55
intros
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    56
  T0 [CPure.intro]: "a \<noteq> b \<Longrightarrow> (ws, zs) \<in> R b \<Longrightarrow> (w # zs, (a # w) # zs) \<in> T a"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    57
  T1 [CPure.intro]: "(ws, zs) \<in> T a \<Longrightarrow> (w # ws, (a # w) # zs) \<in> T a"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    58
  T2 [CPure.intro]: "a \<noteq> b \<Longrightarrow> (ws, zs) \<in> T a \<Longrightarrow> (ws, (b # w) # zs) \<in> T a"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    59
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    60
consts
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    61
  bar :: "letter list list set"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    62
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    63
inductive bar
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    64
intros
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    65
  bar1 [CPure.intro]: "ws \<in> good \<Longrightarrow> ws \<in> bar"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    66
  bar2 [CPure.intro]: "(\<forall>w. w # ws \<in> bar) \<Longrightarrow> ws \<in> bar"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    67
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    68
theorem prop1: "([] # ws) \<in> bar" by rules
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    69
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    70
theorem lemma1: "ws \<in> L as \<Longrightarrow> ws \<in> L (a # as)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    71
  by (erule L.induct, rules+)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    72
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    73
theorem lemma2' [rule_format]: "(vs, ws) \<in> R a \<Longrightarrow> vs \<in> L as \<longrightarrow> ws \<in> L (a # as)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    74
  apply (erule R.induct)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    75
  apply (rule impI)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    76
  apply (erule L.elims)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    77
  apply simp+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    78
  apply (rule impI)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    79
  apply (erule L.elims)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    80
  apply simp_all
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    81
  apply (rule L0)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    82
  apply (erule emb2)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    83
  apply (erule L1)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    84
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    85
 
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    86
theorem lemma2 [rule_format]: "(vs, ws) \<in> R a \<Longrightarrow> vs \<in> good \<longrightarrow> ws \<in> good"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    87
  apply (erule R.induct)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    88
  apply rules
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    89
  apply (rule impI)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    90
  apply (erule good.elims)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    91
  apply simp_all
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    92
  apply (rule good0)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    93
  apply (erule lemma2')
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    94
  apply assumption
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    95
  apply (erule good1)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    96
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    97
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    98
theorem lemma3' [rule_format]:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    99
    "(vs, ws) \<in> T a \<Longrightarrow> vs \<in> L as \<longrightarrow> ws \<in> L (a # as)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   100
  apply (erule T.induct)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   101
  apply (rule impI)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   102
  apply (erule L.elims)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   103
  apply simp_all
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   104
  apply (rule L0)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   105
  apply (erule emb2)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   106
  apply (rule L1)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   107
  apply (erule lemma1)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   108
  apply (rule impI)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   109
  apply (erule L.elims)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   110
  apply simp_all
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   111
  apply rules+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   112
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   113
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   114
theorem lemma3 [rule_format]: "(ws, zs) \<in> T a \<Longrightarrow> ws \<in> good \<longrightarrow> zs \<in> good"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   115
  apply (erule T.induct)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   116
  apply (rule impI)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   117
  apply (erule good.elims)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   118
  apply simp_all
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   119
  apply (rule good0)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   120
  apply (erule lemma1)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   121
  apply (erule good1)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   122
  apply (rule impI)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   123
  apply (erule good.elims)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   124
  apply simp_all
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   125
  apply (rule good0)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   126
  apply (erule lemma3')
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   127
  apply rules+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   128
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   129
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   130
theorem letter_cases:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   131
  "(a::letter) \<noteq> b \<Longrightarrow> (x = a \<Longrightarrow> P) \<Longrightarrow> (x = b \<Longrightarrow> P) \<Longrightarrow> P"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   132
  apply (case_tac a)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   133
  apply (case_tac b)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   134
  apply (case_tac x, simp, simp)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   135
  apply (case_tac x, simp, simp)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   136
  apply (case_tac b)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   137
  apply (case_tac x, simp, simp)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   138
  apply (case_tac x, simp, simp)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   139
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   140
  
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   141
theorem prop2:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   142
  "xs \<in> bar \<Longrightarrow> \<forall>ys. ys \<in> bar \<longrightarrow>
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   143
     (\<forall>a b zs. a \<noteq> b \<longrightarrow> (xs, zs) \<in> T a \<longrightarrow> (ys, zs) \<in> T b \<longrightarrow> zs \<in> bar)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   144
  apply (erule bar.induct)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   145
  apply (rule allI impI)+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   146
  apply (rule bar1)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   147
  apply (rule lemma3)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   148
  apply assumption+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   149
  apply (rule allI, rule impI)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   150
  apply (erule bar.induct)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   151
  apply (rule allI impI)+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   152
  apply (rule bar1)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   153
  apply (rule lemma3, assumption, assumption)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   154
  apply (rule allI impI)+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   155
  apply (rule bar2)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   156
  apply (rule allI)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   157
  apply (induct_tac w)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   158
  apply (rule prop1)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   159
  apply (rule_tac x=aa in letter_cases, assumption)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   160
  apply hypsubst
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   161
  apply (erule_tac x=list in allE)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   162
  apply (erule conjE)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   163
  apply (erule_tac x=wsa in allE, erule impE)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   164
  apply (rule bar2)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   165
  apply rules
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   166
  apply (erule allE, erule allE, erule_tac x="(a # list) # zs" in allE,
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   167
    erule impE, assumption)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   168
  apply (erule impE)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   169
  apply (rule T1)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   170
  apply assumption
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   171
  apply (erule mp)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   172
  apply (rule T2)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   173
  apply (erule not_sym)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   174
  apply assumption
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   175
  apply hypsubst
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   176
  apply (rotate_tac 1)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   177
  apply (erule_tac x=list in allE)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   178
  apply (erule conjE)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   179
  apply (erule allE, erule allE, erule_tac x="(b # list) # zs" in allE,
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   180
    erule impE)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   181
  apply assumption
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   182
  apply (erule impE)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   183
  apply (rule T2)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   184
  apply assumption
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   185
  apply assumption
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   186
  apply (erule mp)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   187
  apply (rule T1)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   188
  apply assumption
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   189
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   190
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   191
theorem lemma4 [rule_format]: "(ws, zs) \<in> R a \<Longrightarrow> ws \<noteq> [] \<longrightarrow> (ws, zs) \<in> T a"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   192
  apply (erule R.induct)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   193
  apply rules
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   194
  apply (rule impI)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   195
  apply (case_tac vs)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   196
  apply (erule R.elims)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   197
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   198
  apply (case_tac a)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   199
  apply (rule_tac b=B in T0)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   200
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   201
  apply (rule R0)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   202
  apply (rule_tac b=A in T0)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   203
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   204
  apply (rule R0)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   205
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   206
  apply (rule T1)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   207
  apply (erule mp)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   208
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   209
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   210
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   211
theorem R_nil: "([], zs) \<in> R a \<Longrightarrow> zs = []"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   212
  by (erule R.elims, simp+)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   213
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   214
theorem letter_eq_dec: "(a::letter) = b \<or> a \<noteq> b"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   215
  apply (case_tac a)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   216
  apply (case_tac b)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   217
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   218
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   219
  apply (case_tac b)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   220
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   221
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   222
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   223
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   224
theorem prop3: "xs \<in> bar \<Longrightarrow> \<forall>zs. (xs, zs) \<in> R a \<longrightarrow> zs \<in> bar"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   225
  apply (erule bar.induct)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   226
  apply (rule allI impI)+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   227
  apply (rule bar1)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   228
  apply (rule lemma2)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   229
  apply assumption+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   230
  apply (rule allI impI)+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   231
  apply (case_tac ws)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   232
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   233
  apply (drule R_nil)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   234
  apply simp_all
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   235
  apply rules
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   236
  apply (rule bar2)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   237
  apply (rule allI)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   238
  apply (induct_tac w)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   239
  apply (rule prop1)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   240
  apply (rule_tac a1=aaa and b1=a in disjE [OF letter_eq_dec])
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   241
  apply rules
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   242
  apply (rule_tac xs="aa # list" and ys="lista # zs" and zs="(aaa # lista) # zs"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   243
    and b=aaa in prop2 [rule_format])
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   244
  apply (rule bar2)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   245
  apply rules
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   246
  apply assumption
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   247
  apply (erule not_sym)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   248
  apply (rule T2)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   249
  apply (erule not_sym)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   250
  apply (erule lemma4)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   251
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   252
  apply (rule T0)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   253
  apply assumption+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   254
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   255
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   256
theorem prop5: "[w] \<in> bar"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   257
  apply (induct_tac w)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   258
  apply (rule prop1)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   259
  apply (rule prop3 [rule_format])
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   260
  apply rules+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   261
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   262
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   263
theorem higman: "[] \<in> bar"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   264
  apply (rule bar2)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   265
  apply (rule allI)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   266
  apply (rule prop5)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   267
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   268
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   269
consts
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   270
  is_prefix :: "'a list \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> bool"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   271
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   272
primrec
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   273
  "is_prefix [] f = True"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   274
  "is_prefix (x # xs) f = (x = f (length xs) \<and> is_prefix xs f)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   275
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   276
theorem good_prefix_lemma:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   277
  "ws \<in> bar \<Longrightarrow> is_prefix ws f \<longrightarrow> (\<exists>vs. is_prefix vs f \<and> vs \<in> good)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   278
  apply (erule bar.induct)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   279
  apply rules
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   280
  apply (rule impI)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   281
  apply (erule_tac x="f (length ws)" in allE)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   282
  apply (erule conjE)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   283
  apply (erule impE)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   284
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   285
  apply assumption
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   286
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   287
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   288
theorem good_prefix: "\<exists>vs. is_prefix vs f \<and> vs \<in> good"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   289
  apply (insert higman)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   290
  apply (drule good_prefix_lemma)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   291
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   292
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   293
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   294
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   295
subsection {* Realizers *}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   296
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   297
subsubsection {* Bar induction *}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   298
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   299
datatype Bar =
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   300
   Good "letter list list"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   301
 | Bar "letter list list" "letter list \<Rightarrow> Bar"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   302
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   303
consts
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   304
  bar_realizes :: "Bar \<Rightarrow> letter list list \<Rightarrow> bool"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   305
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   306
primrec
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   307
  "bar_realizes (Good ws') ws = (ws = ws' \<and> ws' \<in> good)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   308
  "bar_realizes (Bar ws' f) ws = (ws = ws' \<and> (\<forall>w. bar_realizes (f w) (w # ws')))"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   309
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   310
theorem Good_realizer: "ws \<in> good \<Longrightarrow> bar_realizes (Good ws) ws"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   311
  by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   312
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   313
theorem Bar_realizer:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   314
  "\<forall>w. bar_realizes (f w) (w # ws) \<Longrightarrow> bar_realizes (Bar ws f) ws"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   315
  by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   316
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   317
consts
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   318
  bar_ind :: "Bar \<Rightarrow> (letter list list \<Rightarrow> 'a) \<Rightarrow>
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   319
    (letter list list \<Rightarrow> (letter list \<Rightarrow> Bar \<times> 'a) \<Rightarrow> 'a) \<Rightarrow> 'a"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   320
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   321
primrec
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   322
  "bar_ind (Good ws) f g = f ws"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   323
  "bar_ind (Bar ws f') f g = g ws (\<lambda>w. (f' w, bar_ind (f' w) f g))"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   324
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   325
theorem Bar_ind_realizer:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   326
  assumes bar: "bar_realizes r x"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   327
  and f: "\<And>ws. ws \<in> good \<Longrightarrow> P (f ws) ws"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   328
  and g: "\<And>ws f. (\<forall>w. bar_realizes (fst (f w)) (w # ws) \<and> P (snd (f w)) (w # ws)) \<Longrightarrow>
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   329
    P (g ws f) ws"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   330
  shows "P (bar_ind r f g) x"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   331
proof -
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   332
  have "\<And>x. bar_realizes r x \<Longrightarrow> P (bar_ind r f g) x"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   333
    apply (induct r)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   334
    apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   335
    apply (rules intro: f)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   336
    apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   337
    apply (rule g)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   338
    apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   339
    done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   340
  thus ?thesis .
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   341
qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   342
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   343
extract_type
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   344
  "typeof bar \<equiv> Type (TYPE(Bar))"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   345
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   346
realizability
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   347
  "realizes r (w : bar) \<equiv> bar_realizes r w"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   348
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   349
realizers
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   350
  bar1: "Good" "Good_realizer"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   351
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   352
  bar2: "Bar" "\<Lambda>ws f. Bar_realizer \<cdot> _ \<cdot> _"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   353
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   354
  bar.induct (P): "\<lambda>x P. bar_ind"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   355
    "\<Lambda>x P r (h1: _) f (h2: _) g.
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   356
       Bar_ind_realizer \<cdot> _ \<cdot> _ \<cdot> (\<lambda>r x. realizes r (P x)) \<cdot> _ \<cdot> _ \<bullet> h1 \<bullet> h2"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   357
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   358
subsubsection {* Lists *}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   359
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   360
theorem list_ind_realizer:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   361
  assumes f: "P f []"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   362
  and g: "\<And>a as r. P r as \<Longrightarrow> P (g a as r) (a # as)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   363
  shows "P (list_rec f g xs) xs"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   364
  apply (induct xs)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   365
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   366
  apply (rule f)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   367
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   368
  apply (rule g)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   369
  apply assumption
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   370
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   371
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   372
realizers
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   373
  list.induct (P): "\<lambda>P xs f g. list_rec f g xs"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   374
    "\<Lambda>P xs f (h: _) g. list_ind_realizer \<cdot> _ \<cdot> _ \<cdot> _ \<cdot> _ \<bullet> h"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   375
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   376
subsubsection {* Letters *}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   377
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   378
theorem letter_exhaust_realizer:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   379
  "(y = A \<Longrightarrow> P r) \<Longrightarrow> (y = B \<Longrightarrow> P s) \<Longrightarrow> P (case y of A \<Rightarrow> r | B \<Rightarrow> s)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   380
  by (case_tac y, simp+)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   381
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   382
realizers
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   383
  letter.exhaust (P): "\<lambda>y P r s. case y of A \<Rightarrow> r | B \<Rightarrow> s"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   384
    "\<Lambda>y P r (h: _) s. letter_exhaust_realizer \<cdot> _ \<cdot> (\<lambda>x. realizes x P) \<cdot> _ \<cdot> _ \<bullet> h"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   385
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   386
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   387
subsection {* Extracting the program *}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   388
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   389
extract good_prefix
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   390
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   391
text {*
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   392
  Program extracted from the proof of @{text good_prefix}:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   393
  @{thm [display] good_prefix_def [no_vars]}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   394
  Corresponding correctness theorem:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   395
  @{thm [display] good_prefix_correctness [no_vars]}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   396
  Program extracted from the proof of @{text good_prefix_lemma}:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   397
  @{thm [display] good_prefix_lemma_def [no_vars]}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   398
  Program extracted from the proof of @{text higman}:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   399
  @{thm [display] higman_def [no_vars]}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   400
  Program extracted from the proof of @{text prop5}:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   401
  @{thm [display] prop5_def [no_vars]}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   402
  Program extracted from the proof of @{text prop1}:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   403
  @{thm [display] prop1_def [no_vars]}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   404
  Program extracted from the proof of @{text prop2}:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   405
  @{thm [display] prop2_def [no_vars]}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   406
  Program extracted from the proof of @{text prop3}:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   407
  @{thm [display] prop3_def [no_vars]}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   408
*}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   409
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   410
generate_code
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   411
  test = good_prefix
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   412
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   413
ML {*
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   414
val a = 16807.0;
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   415
val m = 2147483647.0;
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   416
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   417
fun nextRand seed =
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   418
  let val t = a*seed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   419
  in  t - m * real (Real.floor(t/m)) end;
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   420
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   421
fun mk_word seed l =
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   422
  let
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   423
    val r = nextRand seed;
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   424
    val i = Real.round (r / m * 10.0);
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   425
  in if i > 7 andalso l > 2 then (r, []) else
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   426
    apsnd (cons (if i mod 2 = 0 then A else B)) (mk_word r (l+1))
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   427
  end;
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   428
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   429
fun f s id0 = mk_word s 0
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   430
  | f s (Suc n) = f (fst (mk_word s 0)) n;
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   431
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   432
val g1 = snd o (f 20000.0);
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   433
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   434
val g2 = snd o (f 50000.0);
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   435
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   436
fun f1 id0 = [A,A]
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   437
  | f1 (Suc id0) = [B]
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   438
  | f1 (Suc (Suc id0)) = [A,B]
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   439
  | f1 _ = [];
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   440
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   441
fun f2 id0 = [A,A]
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   442
  | f2 (Suc id0) = [B]
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   443
  | f2 (Suc (Suc id0)) = [B,A]
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   444
  | f2 _ = [];
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   445
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   446
val xs1 = test g1;
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   447
val xs2 = test g2;
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   448
val xs3 = test f1;
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   449
val xs4 = test f2;
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   450
*}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   451
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   452
end