ProgTutorial/Package/Ind_Code.thy
author Christian Urban <urbanc@in.tum.de>
Sun, 02 Aug 2009 08:44:41 +0200
changeset 299 d0b81d6e1b28
parent 295 24c68350d059
child 301 2728e8daebc0
permissions -rw-r--r--
updated to Isabelle changes and merged sections in the FirstSteps chapter
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
91
667a0943c40b added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Ind_Code
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
     2
imports "../Base" "../FirstSteps" Ind_General_Scheme
91
667a0943c40b added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
667a0943c40b added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
     5
section {* The Gory Details\label{sec:code} *} 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
     6
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
     7
text {*
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
     8
  As mentioned before the code falls roughly into three parts: the code that deals
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
     9
  with the definitions, with the induction principles and with the introduction
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    10
  rules. In addition there are some administrative functions that string everything 
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    11
  together.
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    12
*}
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
    13
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
    14
subsection {* Definitions *}
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
    15
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
    16
text {*
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    17
  We first have to produce for each predicate the user specifies an appropriate
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    18
  definition, whose general form is
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
    19
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
    20
  @{text [display] "pred \<equiv> \<lambda>zs. \<forall>preds. orules \<longrightarrow> pred zs"}
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
    21
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
    22
  and then ``register'' the definition inside a local theory. 
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    23
  To do the latter, we use the following wrapper for the function
256
1fb8d62c88a0 added some first index-information
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
    24
  @{ML [index] define in LocalTheory}. The wrapper takes a predicate name, a syntax
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
    25
  annotation and a term representing the right-hand side of the definition.
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
    26
*}
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
    27
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    28
ML %linenosgray{*fun make_defn ((predname, mx), trm) lthy =
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    29
let 
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    30
  val arg = ((predname, mx), (Attrib.empty_binding, trm))
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
    31
  val ((_, (_ , thm)), lthy') = LocalTheory.define Thm.internalK arg lthy
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    32
in 
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
    33
  (thm, lthy') 
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    34
end*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    35
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    36
text {*
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    37
  It returns the definition (as a theorem) and the local theory in which the
256
1fb8d62c88a0 added some first index-information
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
    38
  definition has been made. In Line 4, @{ML [index] internalK in Thm} is a flag
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
    39
  attached to the theorem (other possibile flags are @{ML [index] definitionK in Thm}
256
1fb8d62c88a0 added some first index-information
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
    40
  and @{ML [index] axiomK in Thm}). These flags just classify theorems and have no
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    41
  significant meaning, except for tools that, for example, find theorems in
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    42
  the theorem database.\footnote{FIXME: put in the section about theorems.} We
256
1fb8d62c88a0 added some first index-information
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
    43
  also use @{ML [index] empty_binding in Attrib} in Line 3, since the definitions for
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    44
  our inductive predicates are not meant to be seen by the user and therefore
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    45
  do not need to have any theorem attributes. A testcase for this function is
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    46
*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    47
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
    48
local_setup %gray {* fn lthy =>
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
    49
let
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    50
  val arg = ((@{binding "My_True"}, NoSyn), @{term True})
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
    51
  val (def, lthy') = make_defn arg lthy 
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
    52
in
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 239
diff changeset
    53
  writeln (string_of_thm_no_vars lthy' def); lthy'
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
    54
end *}
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    55
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    56
text {*
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    57
  which introduces the definition @{thm My_True_def} and then prints it out. 
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
    58
  Since we are testing the function inside \isacommand{local\_setup}, i.e., make
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    59
  actual changes to the ambient theory, we can query the definition with the usual
183
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
    60
  command \isacommand{thm}:
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    61
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    62
  \begin{isabelle}
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    63
  \isacommand{thm}~@{thm [source] "My_True_def"}\\
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    64
  @{text ">"}~@{thm "My_True_def"}
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    65
  \end{isabelle}
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    66
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
    67
  The next two functions construct the right-hand sides of the definitions, 
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    68
  which are terms whose general form is:
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    69
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    70
  @{text [display] "\<lambda>zs. \<forall>preds. orules \<longrightarrow> pred zs"}
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    71
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    72
  When constructing these terms, the variables @{text "zs"} need to be chosen so 
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    73
  that they do not occur in the @{text orules} and also be distinct from the 
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    74
  @{text "preds"}.
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
    75
183
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
    76
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    77
  The first function, named @{text defn_aux}, constructs the term for one
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    78
  particular predicate (the argument @{text "pred"} in the code below). The
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    79
  number of arguments of this predicate is determined by the number of
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    80
  argument types given in @{text "arg_tys"}. The other arguments of the
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    81
  function are the @{text orules} and all the @{text "preds"}.
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    82
*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    83
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
    84
ML %linenosgray{*fun defn_aux lthy orules preds (pred, arg_tys) =
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    85
let 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    86
  fun mk_all x P = HOLogic.all_const (fastype_of x) $ lambda x P
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    87
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    88
  val fresh_args = 
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
    89
        arg_tys 
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    90
        |> map (pair "z")
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    91
        |> Variable.variant_frees lthy (preds @ orules) 
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    92
        |> map Free
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    93
in
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    94
  list_comb (pred, fresh_args)
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    95
  |> fold_rev (curry HOLogic.mk_imp) orules
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    96
  |> fold_rev mk_all preds
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    97
  |> fold_rev lambda fresh_args 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    98
end*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    99
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   100
text {*
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   101
  The function @{text mk_all} in Line 3 is just a helper function for constructing 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   102
  universal quantifications. The code in Lines 5 to 9 produces the fresh @{text
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   103
  "zs"}. For this it pairs every argument type with the string
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   104
  @{text [quotes] "z"} (Line 7); then generates variants for all these strings
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   105
  so that they are unique w.r.t.~to the predicates and @{text "orules"} (Line 8);
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   106
  in Line 9 it generates the corresponding variable terms for the unique
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   107
  strings.
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   108
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   109
  The unique variables are applied to the predicate in Line 11 using the
183
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   110
  function @{ML list_comb}; then the @{text orules} are prefixed (Line 12); in
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   111
  Line 13 we quantify over all predicates; and in line 14 we just abstract
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   112
  over all the @{text "zs"}, i.e., the fresh arguments of the
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   113
  predicate. A testcase for this function is
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   114
*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   115
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   116
local_setup %gray {* fn lthy =>
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   117
let
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   118
  val def = defn_aux lthy eo_orules eo_preds (e_pred, e_arg_tys)
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   119
in
239
b63c72776f03 replaced "warning" with "writeln"
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
   120
  writeln (Syntax.string_of_term lthy def); lthy
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   121
end *}
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   122
91
667a0943c40b added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
text {*
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   124
  where we use the shorthands defined in Figure~\ref{fig:shorthands}.
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   125
  The testcase calls @{ML defn_aux} for the predicate @{text "even"} and prints
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   126
  out the generated definition. So we obtain as printout 
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   127
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   128
  @{text [display] 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   129
"\<lambda>z. \<forall>even odd. (even 0) \<longrightarrow> (\<forall>n. odd n \<longrightarrow> even (Suc n)) 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   130
                         \<longrightarrow> (\<forall>n. even n \<longrightarrow> odd (Suc n)) \<longrightarrow> even z"}
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   131
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   132
  If we try out the function with the rules for freshness
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   133
*}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   134
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   135
local_setup %gray {* fn lthy =>
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   136
let
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   137
  val def = defn_aux lthy fresh_orules 
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   138
                             [fresh_pred] (fresh_pred, fresh_arg_tys)
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   139
in
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   140
  writeln (Syntax.string_of_term lthy def); lthy
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   141
end *}
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   142
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   143
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   144
text {*
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   145
  we obtain
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   146
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   147
  @{term [display] 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   148
"\<lambda>z za. \<forall>fresh. (\<forall>a b. \<not> a = b \<longrightarrow> fresh a (Var b)) \<longrightarrow>
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   149
               (\<forall>a s t. fresh a t \<longrightarrow> fresh a s \<longrightarrow> fresh a (App t s)) \<longrightarrow>
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   150
                (\<forall>a t. fresh a (Lam a t)) \<longrightarrow>
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   151
                (\<forall>a b t. \<not> a = b \<longrightarrow> fresh a t \<longrightarrow> fresh a (Lam b t)) \<longrightarrow> fresh z za"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   152
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   153
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   154
  The second function, named @{text defns}, has to iterate the function
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   155
  @{ML defn_aux} over all predicates. The argument @{text "preds"} is again
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   156
  the list of predicates as @{ML_type term}s; the argument @{text
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   157
  "prednames"} is the list of binding names of the predicates; @{text mxs} 
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   158
  are the list of syntax, or mixfix, annotations for the predicates; 
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   159
  @{text "arg_tyss"} is the list of argument-type-lists.
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   160
*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   161
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   162
ML %linenosgray{*fun defns rules preds prednames mxs arg_typss lthy =
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   163
let
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   164
  val thy = ProofContext.theory_of lthy
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   165
  val orules = map (ObjectLogic.atomize_term thy) rules
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   166
  val defs = map (defn_aux lthy orules preds) (preds ~~ arg_typss) 
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   167
in
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   168
  fold_map make_defn (prednames ~~ mxs ~~ defs) lthy
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   169
end*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   170
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   171
text {*
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   172
  The user will state the introduction rules using meta-implications and
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   173
  meta-quanti\-fications. In Line 4, we transform these introduction rules
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   174
  into the object logic (since definitions cannot be stated with
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   175
  meta-connectives). To do this transformation we have to obtain the theory
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   176
  behind the local theory using the function @{ML [index] theory_of in ProofContext} 
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   177
  (Line 3); with this theory we can use the function
256
1fb8d62c88a0 added some first index-information
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   178
  @{ML [index] atomize_term in ObjectLogic} to make the transformation (Line 4). The call
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   179
  to @{ML defn_aux} in Line 5 produces all right-hand sides of the
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   180
  definitions. The actual definitions are then made in Line 7.  The result of
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   181
  the function is a list of theorems and a local theory (the theorems are
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   182
  registered with the local theory). A testcase for this function is
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   183
*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   184
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   185
local_setup %gray {* fn lthy =>
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   186
let
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   187
  val (defs, lthy') = 
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   188
    defns eo_rules eo_preds eo_prednames eo_mxs eo_arg_tyss lthy
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   189
in
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 239
diff changeset
   190
  writeln (string_of_thms_no_vars lthy' defs); lthy
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   191
end *}
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   192
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   193
text {*
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   194
  where we feed into the function all parameters corresponding to
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   195
  the @{text even}/@{text odd} example. The definitions we obtain
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   196
  are:
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   197
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   198
  @{text [display, break]
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   199
"even \<equiv> \<lambda>z. \<forall>even odd. (even 0) \<longrightarrow> (\<forall>n. odd n \<longrightarrow> even (Suc n))  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   200
                                \<longrightarrow> (\<forall>n. even n \<longrightarrow> odd (Suc n)) \<longrightarrow> even z,
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   201
odd \<equiv> \<lambda>z. \<forall>even odd. (even 0) \<longrightarrow> (\<forall>n. odd n \<longrightarrow> even (Suc n)) 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   202
                               \<longrightarrow> (\<forall>n. even n \<longrightarrow> odd (Suc n)) \<longrightarrow> odd z"}
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   203
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   204
  Note that in the testcase we return the local theory @{text lthy} 
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   205
  (not the modified @{text lthy'}). As a result the test case has no effect
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   206
  on the ambient theory. The reason is that if we introduce the
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   207
  definition again, we pollute the name space with two versions of 
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   208
  @{text "even"} and @{text "odd"}. We want to avoid this here.
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   209
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   210
  This completes the code for introducing the definitions. Next we deal with
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   211
  the induction principles. 
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   212
*}
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   213
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   214
subsection {* Induction Principles *}
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   215
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   216
text {*
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   217
  Recall that the manual proof for the induction principle 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   218
  of @{text "even"} was:
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   219
*}
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   220
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   221
lemma manual_ind_prin_even: 
209
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
   222
assumes prem: "even z"
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
   223
shows "P 0 \<Longrightarrow> (\<And>m. Q m \<Longrightarrow> P (Suc m)) \<Longrightarrow> (\<And>m. P m \<Longrightarrow> Q (Suc m)) \<Longrightarrow> P z"
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   224
apply(atomize (full))
209
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
   225
apply(cut_tac prem)
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   226
apply(unfold even_def)
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   227
apply(drule spec[where x=P])
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   228
apply(drule spec[where x=Q])
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   229
apply(assumption)
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   230
done
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   231
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   232
text {* 
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   233
  The code for automating such induction principles has to accomplish two tasks: 
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   234
  constructing the induction principles from the given introduction
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   235
  rules and then automatically generating proofs for them using a tactic. 
183
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   236
  
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   237
  The tactic will use the following helper function for instantiating universal 
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   238
  quantifiers. 
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   239
*}
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   240
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   241
ML{*fun inst_spec ctrm = 
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   242
  Drule.instantiate' [SOME (ctyp_of_term ctrm)] [NONE, SOME ctrm] @{thm spec}*}
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   243
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   244
text {*
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   245
  This helper function instantiates the @{text "?x"} in the theorem 
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   246
  @{thm spec} with a given @{ML_type cterm}. We call this helper function
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   247
  in the following tactic.\label{fun:instspectac}.
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   248
*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   249
183
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   250
ML{*fun inst_spec_tac ctrms = 
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   251
  EVERY' (map (dtac o inst_spec) ctrms)*}
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   252
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   253
text {*
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   254
  This tactic expects a list of @{ML_type cterm}s. It allows us in the 
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   255
  proof below to instantiate the three quantifiers in the assumption. 
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   256
*}
183
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   257
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   258
lemma 
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   259
fixes P::"nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool"
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   260
shows "\<forall>x y z. P x y z \<Longrightarrow> True"
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   261
apply (tactic {* 
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   262
  inst_spec_tac [@{cterm "a::nat"},@{cterm "b::nat"},@{cterm "c::nat"}] 1 *})
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   263
txt {* 
183
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   264
  We obtain the goal state
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   265
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   266
  \begin{minipage}{\textwidth}
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   267
  @{subgoals} 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   268
  \end{minipage}*}
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   269
(*<*)oops(*>*)
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   270
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   271
text {*
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   272
  The complete tactic for proving the induction principles can now
183
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   273
  be implemented as follows:
163
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   274
*}
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   275
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   276
ML %linenosgray{*fun ind_tac defs prem insts =
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   277
  EVERY1 [ObjectLogic.full_atomize_tac,
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   278
          cut_facts_tac prem,
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   279
          K (rewrite_goals_tac defs),
183
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   280
          inst_spec_tac insts,
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   281
          assume_tac]*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   282
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   283
text {*
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   284
  We have to give it as arguments the definitions, the premise (a list of
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   285
  formulae) and the instantiations. The premise is @{text "even n"} in lemma
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   286
  @{thm [source] manual_ind_prin_even} shown above; in our code it will always be a list
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   287
  consisting of a single formula. Compare this tactic with the manual proof
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   288
  for the lemma @{thm [source] manual_ind_prin_even}: as you can see there is
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   289
  almost a one-to-one correspondence between the \isacommand{apply}-script and
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   290
  the @{ML ind_tac}. We first rewrite the goal to use only object connectives (Line 2),
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   291
  "cut in" the premise (Line 3), unfold the definitions (Line 4), instantiate
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   292
  the assumptions of the goal (Line 5) and then conclude with @{ML assume_tac}.
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   293
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   294
  Two testcases for this tactic are:
183
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   295
*}
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   296
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   297
lemma automatic_ind_prin_even:
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   298
assumes prem: "even z"
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   299
shows "P 0 \<Longrightarrow> (\<And>m. Q m \<Longrightarrow> P (Suc m)) \<Longrightarrow> (\<And>m. P m \<Longrightarrow> Q (Suc m)) \<Longrightarrow> P z"
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   300
by (tactic {* ind_tac eo_defs @{thms prem} 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   301
                    [@{cterm "P::nat\<Rightarrow>bool"}, @{cterm "Q::nat\<Rightarrow>bool"}] *})
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   302
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   303
lemma automatic_ind_prin_fresh:
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   304
assumes prem: "fresh z za" 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   305
shows "(\<And>a b. a \<noteq> b \<Longrightarrow> P a (Var b)) \<Longrightarrow> 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   306
        (\<And>a t s. \<lbrakk>P a t; P a s\<rbrakk> \<Longrightarrow> P a (App t s)) \<Longrightarrow>
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   307
        (\<And>a t. P a (Lam a t)) \<Longrightarrow> 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   308
        (\<And>a b t. \<lbrakk>a \<noteq> b; P a t\<rbrakk> \<Longrightarrow> P a (Lam b t)) \<Longrightarrow> P z za"
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   309
by (tactic {* ind_tac @{thms fresh_def} @{thms prem} 
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   310
                          [@{cterm "P::string\<Rightarrow>trm\<Rightarrow>bool"}] *})
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   311
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   312
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   313
text {*
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   314
  While the tactic for proving the induction principles is relatively simple,
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   315
  it will be a bit more work to construct the goals from the introduction rules
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   316
  the user provides.  Therefore let us have a closer look at the first 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   317
  proved theorem:
209
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
   318
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
   319
  \begin{isabelle}
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   320
  \isacommand{thm}~@{thm [source] automatic_ind_prin_even}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   321
  @{text "> "}~@{thm automatic_ind_prin_even}
209
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
   322
  \end{isabelle}
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
   323
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   324
  The variables @{text "z"}, @{text "P"} and @{text "Q"} are schematic
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   325
  variables (since they are not quantified in the lemma). These 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   326
  variables must be schematic, otherwise they cannot be instantiated 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   327
  by the user. To generate these schematic variables we use a common trick
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   328
  in Isabelle programming: we first declare them as \emph{free}, 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   329
  \emph{but fixed}, and then use the infrastructure to turn them into 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   330
  schematic variables.
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   331
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   332
  In general we have to construct for each predicate @{text "pred"} a goal 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   333
  of the form
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   334
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   335
  @{text [display] 
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   336
  "pred ?zs \<Longrightarrow> rules[preds := ?Ps] \<Longrightarrow> ?P ?zs"}
180
9c25418db6f0 added a recipy about SAT solvers
Christian Urban <urbanc@in.tum.de>
parents: 179
diff changeset
   337
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   338
  where the predicates @{text preds} are replaced in @{text rules} by new 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   339
  distinct variables @{text "?Ps"}. We also need to generate fresh arguments 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   340
  @{text "?zs"} for the predicate  @{text "pred"} and the @{text "?P"} in 
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   341
  the conclusion. 
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   342
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   343
  We generate these goals in two steps. The first function, named @{text prove_ind}, 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   344
  expects that the introduction rules are already appropriately substituted. The argument
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   345
  @{text "srules"} stands for these substituted rules; @{text cnewpreds} are
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   346
  the certified terms coresponding to the variables @{text "?Ps"}; @{text
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   347
  "pred"} is the predicate for which we prove the induction principle;
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   348
  @{text "newpred"} is its replacement and @{text "arg_tys"} are the argument
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   349
  types of this predicate.
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   350
*}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   351
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   352
ML %linenosgray{*fun prove_ind lthy defs srules cnewpreds ((pred, newpred), arg_tys) =
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   353
let
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   354
  val zs = replicate (length arg_tys) "z"
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   355
  val (newargnames, lthy') = Variable.variant_fixes zs lthy;
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   356
  val newargs = map Free (newargnames ~~ arg_tys)
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   357
  
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   358
  val prem = HOLogic.mk_Trueprop (list_comb (pred, newargs))
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   359
  val goal = Logic.list_implies 
183
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   360
         (srules, HOLogic.mk_Trueprop (list_comb (newpred, newargs)))
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   361
in
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   362
  Goal.prove lthy' [] [prem] goal
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   363
      (fn {prems, ...} => ind_tac defs prems cnewpreds)
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   364
  |> singleton (ProofContext.export lthy' lthy)
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   365
end *}
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   366
180
9c25418db6f0 added a recipy about SAT solvers
Christian Urban <urbanc@in.tum.de>
parents: 179
diff changeset
   367
text {* 
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   368
  In Line 3 we produce names @{text "zs"} for each type in the 
183
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   369
  argument type list. Line 4 makes these names unique and declares them as 
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   370
  free, but fixed, variables in the local theory @{text "lthy'"}. 
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   371
  That means they are not schematic variables (yet).
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   372
  In Line 5 we construct the terms corresponding to these variables. 
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   373
  The variables are applied to the predicate in Line 7 (this corresponds
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   374
  to the first premise @{text "pred zs"} of the induction principle). 
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   375
  In Line 8 and 9, we first construct the term  @{text "P zs"} 
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   376
  and then add the (substituted) introduction rules as preconditions. In 
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   377
  case that no introduction rules are given, the conclusion of this 
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   378
  implication needs to be wrapped inside a @{term Trueprop}, otherwise 
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   379
  the Isabelle's goal mechanism will fail.\footnote{FIXME: check with 
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   380
  Stefan...is this so?} 
180
9c25418db6f0 added a recipy about SAT solvers
Christian Urban <urbanc@in.tum.de>
parents: 179
diff changeset
   381
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   382
  In Line 11 we set up the goal to be proved using the function @{ML [index]
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   383
  prove in Goal}; in the next line we call the tactic for proving the
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   384
  induction principle. As mentioned before, this tactic expects the
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   385
  definitions, the premise and the (certified) predicates with which the
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   386
  introduction rules have been substituted. The code in these two lines will
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   387
  return a theorem. However, it is a theorem proved inside the local theory
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   388
  @{text "lthy'"}, where the variables @{text "zs"} are free, but fixed (see
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   389
  Line 4). By exporting this theorem from @{text "lthy'"} (which contains the
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   390
  @{text "zs"} as free variables) to @{text "lthy"} (which does not), we
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   391
  obtain the desired schematic variables @{text "?zs"}.  A testcase for this
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   392
  function is
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   393
*}
180
9c25418db6f0 added a recipy about SAT solvers
Christian Urban <urbanc@in.tum.de>
parents: 179
diff changeset
   394
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   395
local_setup %gray {* fn lthy =>
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   396
let
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   397
  val newpreds = [@{term "P::nat \<Rightarrow> bool"}, @{term "Q::nat \<Rightarrow> bool"}]
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   398
  val cnewpreds = [@{cterm "P::nat \<Rightarrow> bool"}, @{cterm "Q::nat \<Rightarrow> bool"}]
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   399
  val newpred = @{term "P::nat \<Rightarrow> bool"}
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   400
  val srules =  map (subst_free (eo_preds ~~ newpreds)) eo_rules
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   401
  val intro = 
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   402
      prove_ind lthy eo_defs srules cnewpreds ((e_pred, newpred), e_arg_tys)
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   403
in
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 239
diff changeset
   404
  writeln (string_of_thm lthy intro); lthy
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   405
end *}
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   406
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   407
text {*
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   408
  This prints out the theorem:
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   409
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   410
  @{text [display]
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   411
  " \<lbrakk>even ?z; P 0; \<And>n. Q n \<Longrightarrow> P (Suc n); \<And>n. P n \<Longrightarrow> Q (Suc n)\<rbrakk> \<Longrightarrow> P ?z"}
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   412
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   413
  The export from @{text lthy'} to @{text lthy} in Line 13 above 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   414
  has correctly turned the free, but fixed, @{text "z"} into a schematic 
209
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
   415
  variable @{text "?z"}; the variables @{text "P"} and @{text "Q"} are not yet
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   416
  schematic. 
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   417
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   418
  We still have to produce the new predicates with which the introduction
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   419
  rules are substituted and iterate @{ML prove_ind} over all
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   420
  predicates. This is what the second function, named @{text inds} does. 
180
9c25418db6f0 added a recipy about SAT solvers
Christian Urban <urbanc@in.tum.de>
parents: 179
diff changeset
   421
*}
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   422
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   423
ML %linenosgray{*fun inds rules defs preds arg_tyss lthy  =
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   424
let
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   425
  val Ps = replicate (length preds) "P"
183
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   426
  val (newprednames, lthy') = Variable.variant_fixes Ps lthy
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   427
  
183
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   428
  val thy = ProofContext.theory_of lthy'
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   429
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   430
  val tyss' = map (fn tys => tys ---> HOLogic.boolT) arg_tyss
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   431
  val newpreds = map Free (newprednames ~~ tyss')
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   432
  val cnewpreds = map (cterm_of thy) newpreds
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   433
  val srules = map (subst_free (preds ~~ newpreds)) rules
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   434
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   435
in
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   436
  map (prove_ind lthy' defs srules cnewpreds) 
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   437
        (preds ~~ newpreds ~~ arg_tyss)
183
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 180
diff changeset
   438
          |> ProofContext.export lthy' lthy
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   439
end*}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   440
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   441
text {*
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   442
  In Line 3, we generate a string @{text [quotes] "P"} for each predicate. 
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   443
  In Line 4, we use the same trick as in the previous function, that is making the 
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   444
  @{text "Ps"} fresh and declaring them as free, but fixed, in
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   445
  the new local theory @{text "lthy'"}. From the local theory we extract
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   446
  the ambient theory in Line 6. We need this theory in order to certify 
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   447
  the new predicates. In Line 8, we construct the types of these new predicates
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   448
  using the given argument types. Next we turn them into terms and subsequently
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   449
  certify them (Line 9 and 10). We can now produce the substituted introduction rules 
256
1fb8d62c88a0 added some first index-information
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   450
  (Line 11) using the function @{ML [index] subst_free}. Line 14 and 15 just iterate 
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   451
  the proofs for all predicates.
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   452
  From this we obtain a list of theorems. Finally we need to export the 
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   453
  fixed variables @{text "Ps"} to obtain the schematic variables @{text "?Ps"} 
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   454
  (Line 16).
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   455
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   456
  A testcase for this function is
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   457
*}
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   458
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   459
local_setup %gray {* fn lthy =>
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   460
let 
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   461
  val ind_thms = inds eo_rules eo_defs eo_preds eo_arg_tyss lthy
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   462
in
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 239
diff changeset
   463
  writeln (string_of_thms lthy ind_thms); lthy
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   464
end *}
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   465
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   466
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   467
text {*
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   468
  which prints out
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   469
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   470
@{text [display]
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   471
"even ?z \<Longrightarrow> ?P1 0 \<Longrightarrow> 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   472
 (\<And>m. ?Pa1 m \<Longrightarrow> ?P1 (Suc m)) \<Longrightarrow> (\<And>m. ?P1 m \<Longrightarrow> ?Pa1 (Suc m)) \<Longrightarrow> ?P1 ?z,
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   473
odd ?z \<Longrightarrow> ?P1 0 \<Longrightarrow>
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   474
 (\<And>m. ?Pa1 m \<Longrightarrow> ?P1 (Suc m)) \<Longrightarrow> (\<And>m. ?P1 m \<Longrightarrow> ?Pa1 (Suc m)) \<Longrightarrow> ?Pa1 ?z"}
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   475
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   476
  Note that now both, the @{text "?Ps"} and the @{text "?zs"}, are schematic
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   477
  variables. The numbers attached to these variables have been introduced by 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   478
  the pretty-printer and are \emph{not} important for the user. 
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   479
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   480
  This completes the code for the induction principles. The final peice
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   481
  of reasoning infrastructure we need are the introduction rules. 
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   482
*}
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   483
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   484
subsection {* Introduction Rules *}
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   485
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   486
text {*
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   487
  Constructing the goals for the introduction rules is easy: they
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   488
  are just the rules given by the user. However, their proofs are 
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   489
  quite a bit more involved than the ones for the induction principles. 
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   490
  To explain the general method, our running example will be
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   491
  the introduction rule
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   492
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   493
  \begin{isabelle}
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   494
  @{prop "\<And>a b t. \<lbrakk>a \<noteq> b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)"}
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   495
  \end{isabelle}
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   496
  
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   497
  about freshness for lambdas. In order to ease somewhat 
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   498
  our work here, we use the following two helper functions.
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   499
*}
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   500
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   501
ML{*val all_elims = fold (fn ct => fn th => th RS inst_spec ct)
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   502
val imp_elims = fold (fn th => fn th' => [th', th] MRS @{thm mp})*}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   503
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   504
text {* 
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   505
  To see what these functions do, let us suppose we have the following three
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   506
  theorems. 
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   507
*}
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   508
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   509
lemma all_elims_test:
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   510
fixes P::"nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool"
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   511
shows "\<forall>x y z. P x y z" sorry
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   512
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   513
lemma imp_elims_test:
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   514
shows "A \<longrightarrow> B \<longrightarrow> C" sorry
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   515
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   516
lemma imp_elims_test':
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   517
shows "A" "B" sorry
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   518
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   519
text {*
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   520
  The function @{ML all_elims} takes a list of (certified) terms and instantiates
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   521
  theorems of the form @{thm [source] all_elims_test}. For example we can instantiate
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   522
  the quantifiers in this theorem with @{term a}, @{term b} and @{term c} as follows:
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   523
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   524
  @{ML_response_fake [display, gray]
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   525
"let
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   526
  val ctrms = [@{cterm \"a::nat\"}, @{cterm \"b::nat\"}, @{cterm \"c::nat\"}]
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   527
  val new_thm = all_elims ctrms @{thm all_elims_test}
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   528
in
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 239
diff changeset
   529
  writeln (string_of_thm_no_vars @{context} new_thm)
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   530
end"
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   531
  "P a b c"}
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   532
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   533
  Note the difference with @{ML inst_spec_tac} from Page~\pageref{fun:instspectac}:
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   534
  @{ML inst_spec_tac} is a tactic which operates on a goal state; in contrast
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   535
  @{ML all_elims} operates on theorems. 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   536
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   537
  Similarly, the function @{ML imp_elims} eliminates preconditions from implications. 
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   538
  For example we can eliminate the preconditions @{text "A"} and @{text "B"} from
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   539
  @{thm [source] imp_elims_test}:
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   540
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   541
  @{ML_response_fake [display, gray]
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   542
"let
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   543
  val res = imp_elims @{thms imp_elims_test'} @{thm imp_elims_test}
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   544
in
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   545
  writeln (string_of_thm_no_vars @{context} res)
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   546
end"
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   547
  "C"}
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   548
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   549
  Now we set up the proof for the introduction rule as follows:
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   550
*}
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   551
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   552
lemma fresh_Lam:
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   553
shows "\<And>a b t. \<lbrakk>a \<noteq> b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)"
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   554
(*<*)oops(*>*)
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   555
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   556
text {*
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   557
  The first step in the proof will be to expand the definitions of freshness
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   558
  and then introduce quantifiers and implications. For this we
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   559
  will use the tactic
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   560
*}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   561
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   562
ML %linenosgray{*fun expand_tac defs =
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   563
  ObjectLogic.rulify_tac 1
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   564
  THEN rewrite_goals_tac defs
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   565
  THEN (REPEAT (resolve_tac [@{thm allI}, @{thm impI}] 1)) *}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   566
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   567
text {*
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   568
  The function in Line 2 ``rulifies'' the lemma.\footnote{FIXME: explain this better} 
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   569
  This will turn out to 
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   570
  be important later on. Applying this tactic in our proof of @{text "fresh_Lem"}
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   571
*}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   572
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   573
(*<*)
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   574
lemma fresh_Lam:
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   575
shows "\<And>a b t. \<lbrakk>a \<noteq> b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)"
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   576
(*>*)
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   577
apply(tactic {* expand_tac @{thms fresh_def} *})
209
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
   578
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
   579
txt {*
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   580
  gives us the goal state
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   581
209
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
   582
  \begin{isabelle}
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   583
  @{subgoals [display]}
209
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
   584
  \end{isabelle}
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   585
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   586
  As you can see, there are parameters (namely @{text "a"}, @{text "b"} and
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   587
  @{text "t"}) which come from the introduction rule and parameters (in the
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   588
  case above only @{text "fresh"}) which come from the universal
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   589
  quantification in the definition @{term "fresh a (App t s)"}.  Similarly,
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   590
  there are assumptions that come from the premises of the rule (namely the
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   591
  first two) and assumptions from the definition of the predicate (assumption
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   592
  three to six). We need to treat these parameters and assumptions
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   593
  differently. In the code below we will therefore separate them into @{text
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   594
  "params1"} and @{text params2}, respectively @{text "prems1"} and @{text
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   595
  "prems2"}. To do this separation, it is best to open a subproof with the
256
1fb8d62c88a0 added some first index-information
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   596
  tactic @{ML [index] SUBPROOF}, since this tactic provides us with the parameters (as
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   597
  list of @{ML_type cterm}s) and the assumptions (as list of @{ML_type thm}s). 
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   598
  The problem with @{ML SUBPROOF}, however, is that it always expects us to 
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   599
  completely discharge the goal (see Section~\ref{sec:simpletacs}). This is 
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   600
  a bit inconvenient for our gradual explanation of the proof here. Therefore
299
d0b81d6e1b28 updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 295
diff changeset
   601
  we use first the function @{ML [index] FOCUS in Subgoal}, which does s
d0b81d6e1b28 updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 295
diff changeset
   602
  ame as @{ML SUBPROOF} 
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   603
  but does not require us to completely discharge the goal. 
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   604
*}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   605
(*<*)oops(*>*)
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   606
text_raw {*
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   607
\begin{figure}[t]
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   608
\begin{minipage}{\textwidth}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   609
\begin{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   610
*}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   611
ML{*fun chop_print params1 params2 prems1 prems2 ctxt =
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   612
let 
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 239
diff changeset
   613
  val s = ["Params1 from the rule:", string_of_cterms ctxt params1] 
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 239
diff changeset
   614
        @ ["Params2 from the predicate:", string_of_cterms ctxt params2] 
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 239
diff changeset
   615
        @ ["Prems1 from the rule:"] @ (map (string_of_thm ctxt) prems1) 
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 239
diff changeset
   616
        @ ["Prems2 from the predicate:"] @ (map (string_of_thm ctxt) prems2) 
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   617
in 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   618
  s |> separate "\n"
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   619
    |> implode
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   620
    |> tracing
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   621
end*}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   622
text_raw{*
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   623
\end{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   624
\end{minipage}
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   625
\caption{A helper function that prints out the parameters and premises that
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   626
  need to be treated differently.\label{fig:chopprint}}
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   627
\end{figure}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   628
*}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   629
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   630
text {*
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   631
  First we calculate the values for @{text "params1/2"} and @{text "prems1/2"}
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   632
  from @{text "params"} and @{text "prems"}, respectively. To better see what is
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   633
  going in our example, we will print out these values using the printing
299
d0b81d6e1b28 updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 295
diff changeset
   634
  function in Figure~\ref{fig:chopprint}. Since @{ML FOCUS in Subgoal} will
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   635
  supply us the @{text "params"} and @{text "prems"} as lists, we can 
256
1fb8d62c88a0 added some first index-information
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   636
  separate them using the function @{ML [index] chop}. 
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   637
*}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   638
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   639
ML %linenosgray{*fun chop_test_tac preds rules =
299
d0b81d6e1b28 updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 295
diff changeset
   640
  Subgoal.FOCUS (fn {params, prems, context, ...} =>
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   641
  let
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   642
    val cparams = map snd params
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   643
    val (params1, params2) = chop (length cparams - length preds) cparams
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   644
    val (prems1, prems2) = chop (length prems - length rules) prems
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   645
  in
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   646
    chop_print params1 params2 prems1 prems2 context; all_tac
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   647
  end) *}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   648
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   649
text {* 
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   650
  For the separation we can rely on the fact that Isabelle deterministically 
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   651
  produces parameters and premises in a goal state. The last parameters
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   652
  that were introduced come from the quantifications in the definitions
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   653
  (see the tactic @{ML expand_tac}).
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   654
  Therefore we only have to subtract in Line 5 the number of predicates (in this
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   655
  case only @{text "1"}) from the lenghts of all parameters. Similarly
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   656
  with the @{text "prems"} in line 6: the last premises in the goal state come from 
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   657
  unfolding the definition of the predicate in the conclusion. So we can 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   658
  just subtract the number of rules from the number of all premises. 
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   659
  To check our calculations we print them out in Line 8 using the
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   660
  function @{ML chop_print} from Figure~\ref{fig:chopprint} and then 
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   661
  just do nothing, that is @{ML all_tac}. Applying this tactic in our example 
209
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
   662
*}
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
   663
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   664
(*<*)
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   665
lemma fresh_Lam:
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   666
shows "\<And>a b t. \<lbrakk>a \<noteq> b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)"
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   667
apply(tactic {* expand_tac @{thms fresh_def} *})
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   668
(*>*)
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   669
apply(tactic {* chop_test_tac [fresh_pred] fresh_rules @{context} 1 *})
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   670
(*<*)oops(*>*)
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   671
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   672
text {*
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   673
  gives
209
17b1512f51af soem slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 208
diff changeset
   674
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   675
  \begin{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   676
  @{text "Params1 from the rule:"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   677
  @{text "a, b, t"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   678
  @{text "Params2 from the predicate:"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   679
  @{text "fresh"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   680
  @{text "Prems1 from the rule:"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   681
  @{term "a \<noteq> b"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   682
  @{text [break]
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   683
"\<forall>fresh.
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   684
      (\<forall>a b. a \<noteq> b \<longrightarrow> fresh a (Var b)) \<longrightarrow>
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   685
      (\<forall>a t s. fresh a t \<longrightarrow> fresh a s \<longrightarrow> fresh a (App t s)) \<longrightarrow>
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   686
      (\<forall>a t. fresh a (Lam a t)) \<longrightarrow> 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   687
      (\<forall>a b t. a \<noteq> b \<longrightarrow> fresh a t \<longrightarrow> fresh a (Lam b t)) \<longrightarrow> fresh a t"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   688
   @{text "Prems2 from the predicate:"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   689
   @{term "\<forall>a b. a \<noteq> b \<longrightarrow> fresh a (Var b)"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   690
   @{term "\<forall>a t s. fresh a t \<longrightarrow> fresh a s \<longrightarrow> fresh a (App t s)"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   691
   @{term "\<forall>a t. fresh a (Lam a t)"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   692
   @{term "\<forall>a b t. a \<noteq> b \<longrightarrow> fresh a t \<longrightarrow> fresh a (Lam b t)"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   693
  \end{isabelle}
208
0634d42bb69f a bit more work on the simple-inductive package
Christian Urban <urbanc@in.tum.de>
parents: 194
diff changeset
   694
192
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   695
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   696
  We now have to select from @{text prems2} the premise 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   697
  that corresponds to the introduction rule we prove, namely:
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   698
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   699
  @{term [display] "\<forall>a b t. a \<noteq> b \<longrightarrow> fresh a t \<longrightarrow> fresh a (Lam a t)"}
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   700
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   701
  To use this premise with @{ML rtac}, we need to instantiate its 
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   702
  quantifiers (with @{text params1}) and transform it into rule 
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   703
  format (using @{ML [index] rulify in ObjectLogic}). So we can modify the 
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   704
  code as follows:
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   705
*}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   706
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   707
ML %linenosgray{*fun apply_prem_tac i preds rules =
299
d0b81d6e1b28 updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 295
diff changeset
   708
  Subgoal.FOCUS (fn {params, prems, context, ...} =>
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   709
  let
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   710
    val cparams = map snd params
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   711
    val (params1, params2) = chop (length cparams - length preds) cparams
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   712
    val (prems1, prems2) = chop (length prems - length rules) prems
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   713
  in
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   714
    rtac (ObjectLogic.rulify (all_elims params1 (nth prems2 i))) 1
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   715
  end) *}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   716
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   717
text {* 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   718
  The argument @{text i} corresponds to the number of the 
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   719
  introduction we want to prove. We will later on let it range
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   720
  from @{text 0} to the number of @{text "rules - 1"}.
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   721
  Below we apply this function with @{text 3}, since 
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   722
  we are proving the fourth introduction rule. 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   723
*}
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   724
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   725
(*<*)
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   726
lemma fresh_Lam:
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   727
shows "\<And>a b t. \<lbrakk>a \<noteq> b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)"
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   728
apply(tactic {* expand_tac @{thms fresh_def} *})
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   729
(*>*)
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   730
apply(tactic {* apply_prem_tac 3 [fresh_pred] fresh_rules @{context} 1 *})
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   731
(*<*)oops(*>*)
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   732
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   733
text {*
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   734
  The goal state we obtain is: 
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   735
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   736
  \begin{isabelle}
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   737
  @{text "1."}~@{text "\<dots> \<Longrightarrow> "}~@{prop "a \<noteq> b"}\\
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   738
  @{text "2."}~@{text "\<dots> \<Longrightarrow> "}~@{prop "fresh a t"}
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   739
  \end{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   740
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   741
  As expected there are two subgoals, where the first comes from the
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   742
  non-recursive premise of the introduction rule and the second comes 
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   743
  from the recursive one. The first goal can be solved immediately 
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   744
  by @{text "prems1"}. The second needs more work. It can be solved 
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   745
  with the other premise in @{text "prems1"}, namely
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   746
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   747
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   748
  @{term [break,display]
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   749
  "\<forall>fresh.
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   750
      (\<forall>a b. a \<noteq> b \<longrightarrow> fresh a (Var b)) \<longrightarrow>
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   751
      (\<forall>a t s. fresh a t \<longrightarrow> fresh a s \<longrightarrow> fresh a (App t s)) \<longrightarrow>
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   752
      (\<forall>a t. fresh a (Lam a t)) \<longrightarrow> 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   753
      (\<forall>a b t. a \<noteq> b \<longrightarrow> fresh a t \<longrightarrow> fresh a (Lam b t)) \<longrightarrow> fresh a t"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   754
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   755
  but we have to instantiate it appropriately. These instantiations
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   756
  come from @{text "params1"} and @{text "prems2"}. We can determine
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   757
  whether we are in the simple or complicated case by checking whether
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   758
  the topmost connective is an @{text "\<forall>"}. The premises in the simple
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   759
  case cannot have such a quantification, since the first step 
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   760
  of @{ML "expand_tac"} was to ``rulify'' the lemma. 
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   761
  The premise of the complicated case must have at least one  @{text "\<forall>"}
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   762
  coming from the quantification over the @{text preds}. So 
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   763
  we can implement the following function
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   764
*}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   765
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   766
ML{*fun prepare_prem params2 prems2 prem =  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   767
  rtac (case prop_of prem of
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   768
           _ $ (Const (@{const_name All}, _) $ _) =>
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   769
                 prem |> all_elims params2 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   770
                      |> imp_elims prems2
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   771
         | _ => prem) *}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   772
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   773
text {* 
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   774
  which either applies the premise outright (the default case) or if 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   775
  it has an outermost universial quantification, instantiates it first 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   776
  with  @{text "params1"} and then @{text "prems1"}. The following 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   777
  tactic will therefore prove the lemma completely.
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   778
*}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   779
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   780
ML{*fun prove_intro_tac i preds rules =
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   781
  SUBPROOF (fn {params, prems, ...} =>
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   782
  let
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   783
    val cparams = map snd params
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   784
    val (params1, params2) = chop (length cparams - length preds) cparams
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   785
    val (prems1, prems2) = chop (length prems - length rules) prems
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   786
  in
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   787
    rtac (ObjectLogic.rulify (all_elims params1 (nth prems2 i))) 1
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   788
    THEN EVERY1 (map (prepare_prem params2 prems2) prems1)
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   789
  end) *}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   790
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   791
text {*
299
d0b81d6e1b28 updated to Isabelle changes and merged sections in the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 295
diff changeset
   792
  Note that the tactic is now @{ML SUBPROOF}, not @{ML FOCUS in Subgoal} anymore. 
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   793
  The full proof of the introduction rule is as follows:
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   794
*}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   795
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   796
lemma fresh_Lam:
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   797
shows "\<And>a b t. \<lbrakk>a \<noteq> b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)"
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   798
apply(tactic {* expand_tac @{thms fresh_def} *})
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   799
apply(tactic {* prove_intro_tac 3 [fresh_pred] fresh_rules @{context} 1 *})
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   800
done
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   801
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   802
text {* 
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   803
  Phew!\ldots  
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   804
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   805
  Unfortunately, not everything is done yet. If you look closely
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   806
  at the general principle outlined for the introduction rules in 
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   807
  Section~\ref{sec:nutshell}, we have  not yet dealt with the case where 
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   808
  recursive premises have preconditions. The introduction rule
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   809
  of the accessible part is such a rule. 
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   810
*}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   811
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   812
lemma accpartI:
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   813
shows "\<And>R x. (\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x"
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   814
apply(tactic {* expand_tac @{thms accpart_def} *})
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   815
apply(tactic {* chop_test_tac [acc_pred] acc_rules @{context} 1 *})
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   816
apply(tactic {* apply_prem_tac 0 [acc_pred] acc_rules @{context} 1 *})
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   817
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   818
txt {*
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   819
  Here @{ML chop_test_tac} prints out the following
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   820
  values for @{text "params1/2"} and @{text "prems1/2"}
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   821
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   822
  \begin{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   823
  @{text "Params1 from the rule:"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   824
  @{text "x"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   825
  @{text "Params2 from the predicate:"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   826
  @{text "P"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   827
  @{text "Prems1 from the rule:"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   828
  @{text "R ?y x \<Longrightarrow> \<forall>P. (\<forall>x. (\<forall>y. R y x \<longrightarrow> P y) \<longrightarrow> P x) \<longrightarrow> P ?y"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   829
  @{text "Prems2 from the predicate:"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   830
  @{term "\<forall>x. (\<forall>y. R y x \<longrightarrow> P y) \<longrightarrow> P x"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   831
  \end{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   832
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   833
  and after application of the introduction rule 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   834
  using @{ML apply_prem_tac}, we are in the goal state
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   835
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   836
  \begin{isabelle}
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   837
  @{text "1."}~@{term "\<And>y. R y x \<Longrightarrow> P y"}
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   838
  \end{isabelle}
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   839
  
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   840
  
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   841
*}(*<*)oops(*>*)
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   842
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   843
text {*
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   844
  In order to make progress, we have to use the precondition
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   845
  @{text "R y x"} (in general there can be many of them). The best way
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   846
  to get a handle on these preconditions is to open up another subproof,
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   847
  since the preconditions will then be bound to @{text prems}. Therfore we
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   848
  modify the function @{ML prepare_prem} as follows
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   849
*}
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   850
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   851
ML %linenosgray{*fun prepare_prem params2 prems2 ctxt prem =  
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   852
  SUBPROOF (fn {prems, ...} =>
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   853
  let
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   854
    val prem' = prems MRS prem
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   855
  in 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   856
    rtac (case prop_of prem' of
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   857
           _ $ (Const (@{const_name All}, _) $ _) =>
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   858
                 prem' |> all_elims params2 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   859
                       |> imp_elims prems2
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   860
         | _ => prem') 1
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   861
  end) ctxt *}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   862
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   863
text {*
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   864
  In Line 4 we use the @{text prems} from the @{ML SUBPROOF} and resolve 
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   865
  them with @{text prem}. In the simple cases, that is where the @{text prem} 
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   866
  comes from a non-recursive premise of the rule, @{text prems} will be 
256
1fb8d62c88a0 added some first index-information
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   867
  just the empty list and the function @{ML [index] MRS} does nothing. Similarly, in the 
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   868
  cases where the recursive premises of the rule do not have preconditions. 
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   869
  In case there are preconditions, then Line 4 discharges them. After
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   870
  that we can proceed as before, i.e., check whether the outermost
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   871
  connective is @{text "\<forall>"}.
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   872
  
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   873
  The function @{ML prove_intro_tac} only needs to be changed so that it
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   874
  gives the context to @{ML prepare_prem} (Line 8). The modified version
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   875
  is below.
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   876
*}
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   877
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   878
ML %linenosgray{*fun prove_intro_tac i preds rules =
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   879
  SUBPROOF (fn {params, prems, context, ...} =>
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   880
  let
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   881
    val cparams = map snd params
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   882
    val (params1, params2) = chop (length cparams - length preds) cparams
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   883
    val (prems1, prems2) = chop (length prems - length rules) prems
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   884
  in
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   885
    rtac (ObjectLogic.rulify (all_elims params1 (nth prems2 i))) 1
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   886
    THEN EVERY1 (map (prepare_prem params2 prems2 context) prems1)
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   887
  end) *}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   888
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   889
text {*
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   890
  With these two functions we can now also prove the introduction
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   891
  rule for the accessible part. 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   892
*}
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   893
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   894
lemma accpartI:
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   895
shows "\<And>R x. (\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x"
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   896
apply(tactic {* expand_tac @{thms accpart_def} *})
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   897
apply(tactic {* prove_intro_tac 0 [acc_pred] acc_rules @{context} 1 *})
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   898
done
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   899
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   900
text {*
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   901
  Finally we need two functions that string everything together. The first
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   902
  function is the tactic that performs the proofs.
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   903
*}
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   904
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   905
ML %linenosgray{*fun intro_tac defs rules preds i ctxt =
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   906
  EVERY1 [ObjectLogic.rulify_tac,
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   907
          K (rewrite_goals_tac defs),
184
c7f04a008c9c some polishing
Christian Urban <urbanc@in.tum.de>
parents: 183
diff changeset
   908
          REPEAT o (resolve_tac [@{thm allI}, @{thm impI}]),
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   909
          prove_intro_tac i preds rules ctxt]*}
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   910
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   911
text {*
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   912
  Lines 2 to 4 in this tactic correspond to the function @{ML expand_tac}. 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   913
  Some testcases for this tactic are:
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   914
*}
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   915
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   916
lemma even0_intro:
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   917
shows "even 0"
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   918
by (tactic {* intro_tac eo_defs eo_rules eo_preds 0 @{context} *})
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   919
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   920
lemma evenS_intro:
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   921
shows "\<And>m. odd m \<Longrightarrow> even (Suc m)"
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   922
by (tactic {* intro_tac eo_defs eo_rules eo_preds 1 @{context} *})
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   923
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   924
lemma fresh_App:
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   925
shows "\<And>a t s. \<lbrakk>fresh a t; fresh a s\<rbrakk> \<Longrightarrow> fresh a (App t s)"
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   926
by (tactic {* 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   927
  intro_tac @{thms fresh_def} fresh_rules [fresh_pred] 1 @{context} *})
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   928
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   929
text {*
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   930
  The second function sets up in Line 4 the goals to be proved (this is easy
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   931
  for the introduction rules since they are exactly the rules 
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   932
  given by the user) and iterates @{ML intro_tac} over all 
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   933
  introduction rules.
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   934
*}
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   935
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   936
ML %linenosgray{*fun intros rules preds defs lthy = 
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   937
let
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   938
  fun intros_aux (i, goal) =
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   939
    Goal.prove lthy [] [] goal
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   940
      (fn {context, ...} => intro_tac defs rules preds i context)
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   941
in
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   942
  map_index intros_aux rules
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   943
end*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   944
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   945
text {*
256
1fb8d62c88a0 added some first index-information
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   946
  The iteration is done with the function @{ML [index] map_index} since we
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   947
  need the introduction rule together with its number (counted from
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   948
  @{text 0}). This completes the code for the functions deriving the
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   949
  reasoning infrastructure. It remains to implement some administrative
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   950
  code that strings everything together.
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   951
*}
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   952
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   953
subsection {* Administrative Functions *}
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   954
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   955
text {* 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   956
  We have produced various theorems (definitions, induction principles and
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   957
  introduction rules), but apart from the definitions, we have not yet 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   958
  registered them with the theorem database. This is what the functions 
256
1fb8d62c88a0 added some first index-information
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   959
  @{ML [index] note in LocalTheory} does. 
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   960
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   961
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   962
  For convenience, we use the following 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   963
  three wrappers this function:
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   964
*}
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   965
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   966
ML{*fun note_many qname ((name, attrs), thms) = 
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   967
  LocalTheory.note Thm.theoremK 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   968
    ((Binding.qualify false qname name, attrs), thms) 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   969
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   970
fun note_single1 qname ((name, attrs), thm) = 
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   971
  note_many qname ((name, attrs), [thm]) 
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   972
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   973
fun note_single2 name attrs (qname, thm) = 
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   974
  note_many (Binding.name_of qname) ((name, attrs), [thm]) *}
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   975
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   976
text {*
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   977
  The function that ``holds everything together'' is @{text "add_inductive"}. 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   978
  Its arguments are the specification of the predicates @{text "pred_specs"} 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   979
  and the introduction rules @{text "rule_spec"}.   
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   980
*}
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   981
186
371e4375c994 made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents: 185
diff changeset
   982
ML %linenosgray{*fun add_inductive pred_specs rule_specs lthy =
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   983
let
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   984
  val mxs = map snd pred_specs
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   985
  val pred_specs' = map fst pred_specs
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   986
  val prednames = map fst pred_specs'
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   987
  val preds = map (fn (p, ty) => Free (Binding.name_of p, ty)) pred_specs'
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   988
  val tyss = map (binder_types o fastype_of) preds   
163
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   989
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   990
  val (namesattrs, rules) = split_list rule_specs    
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   991
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   992
  val (defs, lthy') = defns rules preds prednames mxs tyss lthy      
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   993
  val ind_prins = inds rules defs preds tyss lthy' 	
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 209
diff changeset
   994
  val intro_rules = intros rules preds defs lthy'
91
667a0943c40b added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   995
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   996
  val mut_name = space_implode "_" (map Binding.name_of prednames)
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   997
  val case_names = map (Binding.name_of o fst) namesattrs
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   998
in
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
   999
  lthy' |> note_many mut_name ((@{binding "intros"}, []), intro_rules) 
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
  1000
        ||>> note_many mut_name ((@{binding "inducts"}, []), ind_prins)
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
  1001
        ||>> fold_map (note_single1 mut_name) (namesattrs ~~ intro_rules)  
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 294
diff changeset
  1002
        ||>> fold_map (note_single2 @{binding "induct"} 
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1003
              [Attrib.internal (K (RuleCases.case_names case_names)),
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1004
               Attrib.internal (K (RuleCases.consumes 1)),
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1005
               Attrib.internal (K (Induct.induct_pred ""))]) 
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
  1006
             (prednames ~~ ind_prins) 
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1007
        |> snd
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
  1008
end*}
91
667a0943c40b added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1009
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1010
text {*
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1011
  In Line 3 the function extracts the syntax annotations from the predicates. 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1012
  Lines 4 to 6 extract the names of the predicates and generate
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1013
  the variables terms (with types) corresponding to the predicates. 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1014
  Line 7 produces the argument types for each predicate. 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1015
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1016
  Line 9 extracts the introduction rules from the specifications
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1017
  and stores also in @{text namesattrs} the names and attributes the
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1018
  user may have attached to these rules.
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1019
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1020
  Line 11 produces the definitions and also registers the definitions
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1021
  in the local theory @{text "lthy'"}. The next two lines produce
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1022
  the induction principles and the introduction rules (all of them
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1023
  as theorems). Both need the local theory @{text lthy'} in which
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1024
  the definitions have been registered.
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1025
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1026
  Lines 15 produces the name that is used to register the introduction
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1027
  rules. It is costum to collect all introduction rules under 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1028
  @{text "string.intros"}, whereby @{text "string"} stands for the 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1029
  @{text [quotes] "_"}-separated list of predicate names (for example
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1030
  @{text "even_odd"}. Also by custom, the case names in intuction 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1031
  proofs correspond to the names of the introduction rules. These
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1032
  are generated in Line 16.
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1033
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
  1034
  Lines 18 and 19 now add to @{text "lthy'"} all the introduction rules 
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
  1035
  und induction principles under the name @{text "mut_name.intros"} and
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
  1036
  @{text "mut_name.inducts"}, respectively (see previous paragraph).
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
  1037
  
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
  1038
  Line 20 add further every introduction rule under its own name
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1039
  (given by the user).\footnote{FIXME: what happens if the user did not give
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
  1040
  any name.} Line 21 registers the induction principles. For this we have
256
1fb8d62c88a0 added some first index-information
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
  1041
  to use some specific attributes. The first @{ML [index] case_names in RuleCases} 
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1042
  corresponds to the case names that are used by Isar to reference the proof
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1043
  obligations in the induction. The second @{ML "consumes 1" in RuleCases}
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1044
  indicates that the first premise of the induction principle (namely
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1045
  the predicate over which the induction proceeds) is eliminated. 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1046
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
  1047
  This completes all the code and fits in with the ``front end'' described
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
  1048
  in Section~\ref{sec:interface}.\footnote{FIXME: Describe @{ML Induct.induct_pred}. 
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
  1049
  Why the mut-name? 
224
647cab4a72c2 finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
  1050
  What does @{ML Binding.qualify} do?}
124
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
  1051
*}
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
  1052
(*<*)end(*>*)