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