CookBook/Package/Ind_Code.thy
author Christian Urban <urbanc@in.tum.de>
Mon, 16 Mar 2009 03:02:56 +0100
changeset 179 75381fa516cd
parent 178 fb8f22dd8ad0
child 180 9c25418db6f0
permissions -rw-r--r--
more work on the simple-induct. chapter
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
91
667a0943c40b added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Ind_Code
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
     2
imports "../Base" "../FirstSteps" Simple_Inductive_Package Ind_Prelims
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
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
     5
section {* Code *}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
     6
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
     7
subsection {* Definitions *}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
     8
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
     9
text {*
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
    10
  @{text [display] "rule ::= \<And>xs. As \<Longrightarrow> (\<And>ys. Bs \<Longrightarrow> pred ss)\<^isup>* \<Longrightarrow> pred ts"}
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
    11
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
    12
  @{text [display] "orule ::= \<forall>xs. As \<longrightarrow> (\<forall>ys. Bs \<longrightarrow> pred ss)\<^isup>* \<longrightarrow> pred ts"}
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
    13
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
    14
  @{text [display] "def ::= 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
    15
  
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
    16
  @{text [display] "ind ::= \<And>zs. pred zs \<Longrightarrow> rules[preds::=Ps] \<Longrightarrow> P 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
    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] "oind ::= \<forall>zs. pred zs \<longrightarrow> orules[preds::=Ps] \<longrightarrow> P 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
  
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
    20
  So we have @{text "pred zs"} and @{text "orules[preds::=Ps]"}; have to show
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
    21
  @{text "P zs"}. Expanding @{text "pred zs"} gives @{text "\<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
    22
  Instantiating the @{text "preds"} with @{text "Ps"} gives
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
  @{text "orules[preds::=Ps] \<longrightarrow> P zs"}. So we can conclude with @{text "P 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
    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
  We have to show @{text "\<forall>xs. As \<longrightarrow> (\<forall>ys. Bs \<longrightarrow> pred ss)\<^isup>* \<longrightarrow> pred ts"};
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
    26
  expanding the defs 
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
    27
  
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
  @{text [display]
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
    29
  "\<forall>xs. As \<longrightarrow> (\<forall>ys. Bs \<longrightarrow> (\<forall>preds. orules \<longrightarrow> pred ss))\<^isup>* \<longrightarrow>  (\<forall>preds. orules \<longrightarrow> pred ts"}
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
    30
  
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
    31
  so we have @{text "As"}, @{text "(\<forall>ys. Bs \<longrightarrow> (\<forall>preds. orules \<longrightarrow> pred ss))\<^isup>*"},
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
    32
  @{text "orules"}; and have to show @{text "pred ts"}
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
    33
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
    34
  the @{text "orules"} are of the form @{text "\<forall>xs. As \<longrightarrow> (\<forall>ys. Bs \<longrightarrow> pred ss)\<^isup>* \<longrightarrow> pred ts"}.
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
    35
  
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
    36
  using the @{text "As"} we ????
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    37
*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    38
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
    39
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
    40
text {*
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
    41
  First we have to produce for each predicate its definitions of the form
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
    42
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
    43
  @{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
    44
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    45
  We use the following wrapper function to make the definition via
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
    46
  @{ML LocalTheory.define}. The function takes a predicate name, a syntax
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
    47
  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
    48
*}
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
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
    50
ML %linenosgray{*fun make_defs ((predname, syn), trm) lthy =
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    51
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
    52
  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
    53
  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
    54
in 
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
    55
  (thm, lthy') 
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    56
end*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    57
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    58
text {*
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    59
  It returns the definition (as theorem) and the local theory in which this definition has 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    60
  been made. In Line 4 @{ML internalK in Thm} is just a flag attached to the 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    61
  theorem (others possibilities are @{ML definitionK in Thm} or @{ML axiomK in Thm}). 
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
    62
  These flags just classify theorems and have no significant meaning, except 
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
    63
  for tools such as finding theorems in the theorem database. We also
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    64
  use @{ML empty_binding in Attrib} in Line 3, since the definition does 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    65
  not need any theorem attributes. Note the definition has not yet been 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    66
  stored in the theorem database. So at the moment we can only refer to it
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    67
  via the return value. A testcase for this functions is
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    68
*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    69
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
    70
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
    71
let
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
    72
  val arg =  ((Binding.name "MyTrue", NoSyn), @{term True})
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
    73
  val (def, lthy') = make_defs arg 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
    74
in
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    75
  warning (str_of_thm 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
    76
end *}
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    77
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    78
text {*
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    79
  which prints out the theorem @{prop "MyTrue \<equiv> True"}. Since we are
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    80
  testing the function inside \isacommand{local\_setup} we have also
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    81
  access to theorem associated with this definition.
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    82
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    83
  \begin{isabelle}
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    84
  \isacommand{thm}~@{text "MyTrue_def"}\\
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    85
  @{text "> MyTrue \<equiv> True"}
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    86
  \end{isabelle}
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    87
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    88
  The next function constructs the term for the definition, namely 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    89
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    90
  @{text [display] "\<lambda>\<^raw:$zs$>. \<forall>preds. orules \<longrightarrow> pred \<^raw:$zs$>"}
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    91
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    92
  The variables @{text "\<^raw:$zs$>"} need to be chosen so to not occur
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    93
  in the @{text orules} and also be distinct from @{text "pred"}. The function
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    94
  constructs the term for one particular predicate @{text "pred"}; the number
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
    95
  of @{text "\<^raw:$zs$>"} is determined by the nunber of types. 
164
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
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
    98
ML %linenosgray{*fun defs_aux lthy orules preds (pred, arg_tys) =
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    99
let 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   100
  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
   101
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   102
  val fresh_args = 
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   103
        arg_tys 
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   104
        |> map (pair "z")
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   105
        |> Variable.variant_frees lthy (preds @ orules) 
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   106
        |> map Free
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   107
in
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   108
  list_comb (pred, fresh_args)
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   109
  |> fold_rev (curry HOLogic.mk_imp) orules
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   110
  |> fold_rev mk_all preds
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   111
  |> fold_rev lambda fresh_args 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   112
end*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   113
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   114
text {*
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   115
  The code in Lines 5 to 9 produce the fresh @{text "\<^raw:$zs$>"} with which the 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   116
  predicate is applied. For this it pairs every type with a string @{text [quotes] "z"} 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   117
  (Line 7); then generates variants for all these strings (names) so that they are 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   118
  unique w.r.t.~to the orules and predicates; in Line 9 it generates the corresponding 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   119
  variable terms for the unique names.
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   120
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   121
  The unique free variables are applied to the predicate (Line 11); then
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   122
  the @{text orules} are prefixed (Line 12); in Line 13 we
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   123
  quantify over all predicates; and in line 14 we just abstract over all
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   124
  the (fresh)  @{text "\<^raw:$zs$>"}, i.e.~the arguments of the predicate.
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   125
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   126
  A testcase for this function is
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   127
*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   128
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   129
local_setup %gray{* fn lthy =>
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   130
let
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   131
  val orules = [@{prop "even 0"},
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   132
                @{prop "\<forall>n::nat. odd n \<longrightarrow> even (Suc n)"},
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   133
                @{prop "\<forall>n::nat. even n \<longrightarrow> odd (Suc n)"}] 
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   134
  val preds = [@{term "even::nat\<Rightarrow>bool"}, @{term "odd::nat\<Rightarrow>bool"}]
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   135
  val pred = @{term "even::nat\<Rightarrow>bool"}
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   136
  val arg_tys = [@{typ "nat"}]
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   137
  val def = defs_aux lthy orules preds (pred, arg_tys)
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   138
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
   139
  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
   140
end *}
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   141
91
667a0943c40b added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
text {*
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   143
  It constructs the term for the predicate @{term "even"}. So we obtain as printout
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   144
  the term
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   145
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   146
  @{text [display] 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   147
"\<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
   148
                         \<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
   149
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   150
  The main function for the definitions now has to just iterate
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   151
  the function @{ML defs_aux} over all predicates. THis is what the
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   152
  next function does. 
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
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   155
ML %linenosgray{*fun definitions rules preds prednames syns arg_typss lthy =
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
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   159
  val defs = map (defs_aux lthy orules preds) (preds ~~ arg_typss) 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   160
in
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   161
  fold_map make_defs (prednames ~~ syns ~~ defs) lthy
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 {*
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   165
  The argument @{text "preds"} is the list of predicates as @{ML_type term}s;
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   166
  the argument @{text "prednames"} is the list of names of the predicates;
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   167
  @{text "arg_tyss"} is the list of argument-type-lists. 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   168
  
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   169
  In line 4 we generate the intro rules in the object logic; for this we have to 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   170
  obtain the theory behind the local theory (Line 3); with this we can
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   171
  call @{ML defs_aux} to generate the terms for the left-hand sides.
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   172
  The actual definitions are made in Line 7.
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   173
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   174
  A testcase for 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
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
   179
  val rules = [@{prop "even 0"},
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   180
               @{prop "\<And>n::nat. odd n \<Longrightarrow> even (Suc n)"},
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   181
               @{prop "\<And>n::nat. even n \<Longrightarrow> odd (Suc n)"}] 
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
   182
  val preds = [@{term "even::nat\<Rightarrow>bool"}, @{term "odd::nat\<Rightarrow>bool"}]
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
   183
  val prednames = [Binding.name "even", Binding.name "odd"] 
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
  val syns = [NoSyn, NoSyn] 
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
   185
  val arg_tyss = [[@{typ "nat"}], [@{typ "nat"}]]
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   186
  val (defs, lthy') = definitions rules preds prednames syns arg_tyss 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
   187
in
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
   188
  warning (str_of_thms lthy' defs); lthy
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   189
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
   190
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   191
text {*
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   192
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   193
  It prints out the two definitions
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
   194
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   195
  @{text [display] 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   196
"even \<equiv> \<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
   197
                                \<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
   198
 odd \<equiv> \<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
   199
                                \<longrightarrow> (\<forall>n. even n \<longrightarrow> odd (Suc n)) \<longrightarrow> odd z"}
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   200
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   201
  This completes the code concerning the definitions. Next comes the code for
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   202
  the induction principles. 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   203
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   204
  Recall the proof for the induction principle for @{term "even"}: 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   205
*}
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   206
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   207
lemma 
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   208
  assumes prems: "even n"
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   209
  shows "P 0 \<Longrightarrow> 
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   210
             (\<And>m. Q m \<Longrightarrow> P (Suc m)) \<Longrightarrow> (\<And>m. P m \<Longrightarrow> Q (Suc m)) \<Longrightarrow> P n"
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   211
apply(atomize (full))
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   212
apply(cut_tac prems)
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   213
apply(unfold even_def)
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   214
apply(drule spec[where x=P])
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   215
apply(drule spec[where x=Q])
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   216
apply(assumption)
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   217
done
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   218
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   219
text {* 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   220
  To automate this proof we need to be able to instantiate universal 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   221
  quantifiers. For this we use the following helper function. It
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   222
  instantiates the @{text "?x"} in @{thm spec} with a @{ML_type cterm}.
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   223
*}
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   224
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   225
ML{*fun inst_spec ctrm = 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   226
 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
   227
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   228
text {*
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   229
  For example we can use it to instantiate an assumption:
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   230
*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   231
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   232
lemma "\<forall>x1 x2 x3. P (x1::nat) (x2::nat) (x3::nat) \<Longrightarrow> True"
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   233
apply (tactic {* 
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   234
  let 
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   235
    val ctrms = [@{cterm "y1::nat"},@{cterm "y2::nat"},@{cterm "y3::nat"}]
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   236
  in 
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   237
    EVERY1 (map (dtac o inst_spec) ctrms)
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   238
  end *})
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   239
txt {* 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   240
  where it produces the goal state
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   241
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   242
  \begin{minipage}{\textwidth}
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   243
  @{subgoals} 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   244
  \end{minipage}*}
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   245
(*<*)oops(*>*)
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   246
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   247
text {*
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   248
  Now the tactic for proving the induction rules can be implemented 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   249
  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
   250
*}
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
   251
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   252
ML{*fun induction_tac defs prems insts =
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   253
  EVERY1 [ObjectLogic.full_atomize_tac,
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   254
          cut_facts_tac prems,
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   255
          K (rewrite_goals_tac defs),
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   256
          EVERY' (map (dtac o inst_spec) insts),
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   257
          assume_tac]*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   258
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   259
text {*
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   260
  We only have to give it as arguments the premises and the instantiations.
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   261
  A testcase for the tactic is
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   262
*}
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   263
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   264
lemma 
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   265
  assumes prems: "even n"
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   266
  shows "P 0 \<Longrightarrow> 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   267
           (\<And>m. Q m \<Longrightarrow> P (Suc m)) \<Longrightarrow> (\<And>m. P m \<Longrightarrow> Q (Suc m)) \<Longrightarrow> P n"
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   268
apply(tactic {* 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   269
  let
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   270
     val defs = [@{thm even_def}, @{thm odd_def}]
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   271
     val insts = [@{cterm "P::nat\<Rightarrow>bool"}, @{cterm "Q::nat\<Rightarrow>bool"}]
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   272
  in 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   273
    induction_tac defs @{thms prems} insts 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   274
  end *})
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   275
done
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   276
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   277
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   278
text {*
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   279
  which indeed proves the lemma. 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   280
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   281
  While the generic proof for the induction principle is relatively simple, 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   282
  it is a bit harder to set up the goals just from the given introduction 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   283
  rules. For this we have to construct
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   284
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   285
  @{text [display] 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   286
  "\<And>\<^raw:$zs$>. pred \<^raw:$zs$> \<Longrightarrow> rules[preds := \<^raw:$Ps$>] \<Longrightarrow> \<^raw:$P$> \<^raw:$zs$>"}
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   287
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   288
  where the given predicates are replaced by new ones written
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   289
  as @{text "\<^raw:$Ps$>"}, and also generate new variables 
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   290
  @{text "\<^raw:$zs$>"}.
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   291
*}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   292
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   293
ML %linenosgray{* fun prove_induction lthy defs rules cnewpreds ((pred, newpred), 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
   294
let
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
   295
  val zs = replicate (length tys) "z"
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   296
  val (newargnames, lthy') = Variable.variant_fixes zs 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
   297
  val newargs = map Free (newargnames ~~ tys)
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
   298
  
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
   299
  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
   300
  val goal = Logic.list_implies 
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
   301
         (rules, HOLogic.mk_Trueprop (list_comb (newpred, 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
   302
in
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   303
  Goal.prove lthy' [] [prem] goal
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
   304
  (fn {prems, ...} => induction_tac defs prems cnewpreds)
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   305
  |> singleton (ProofContext.export lthy' lthy)
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   306
end *}
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   307
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   308
text {* *}
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   309
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   310
ML %linenosgray{*fun inductions rules defs preds tyss lthy1  =
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   311
let
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   312
  val Ps = replicate (length preds) "P"
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   313
  val (newprednames, lthy2) = Variable.variant_fixes Ps lthy1
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   314
  
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   315
  val thy = ProofContext.theory_of lthy2
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   316
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   317
  val tyss' = map (fn tys => tys ---> HOLogic.boolT) tyss
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   318
  val newpreds = map Free (newprednames ~~ tyss')
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   319
  val cnewpreds = map (cterm_of thy) newpreds
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   320
  val rules' = map (subst_free (preds ~~ newpreds)) rules
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   321
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   322
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
   323
  map (prove_induction lthy2 defs rules' cnewpreds) 
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
   324
        (preds ~~ newpreds ~~ tyss)
179
75381fa516cd more work on the simple-induct. chapter
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   325
          |> ProofContext.export lthy2 lthy1
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   326
end*}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   327
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   328
ML {*
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   329
  let 
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   330
    val rules = [@{prop "even (0::nat)"},
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   331
                 @{prop "\<And>n::nat. odd n \<Longrightarrow> even (Suc n)"},
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   332
                 @{prop "\<And>n::nat. even n \<Longrightarrow> odd (Suc n)"}] 
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   333
    val defs = [@{thm even_def}, @{thm odd_def}]
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   334
    val preds = [@{term "even::nat\<Rightarrow>bool"}, @{term "odd::nat\<Rightarrow>bool"}]
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   335
    val tyss = [[@{typ "nat"}], [@{typ "nat"}]]
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   336
  in
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   337
    inductions rules defs preds tyss @{context}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   338
  end
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   339
*}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   340
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   341
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   342
subsection {* Introduction Rules *}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   343
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   344
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
   345
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
   346
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   347
ML{*fun subproof2 prem params2 prems2 =  
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   348
 SUBPROOF (fn {prems, ...} =>
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   349
   let
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   350
     val prem' = prems MRS prem;
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   351
     val prem'' = 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   352
       case prop_of prem' of
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   353
           _ $ (Const (@{const_name All}, _) $ _) =>
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   354
             prem' |> all_elims params2 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   355
                   |> imp_elims prems2
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   356
         | _ => prem';
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   357
   in 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   358
     rtac prem'' 1 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   359
   end)*}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   360
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   361
ML{*fun subproof1 rules preds i = 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   362
 SUBPROOF (fn {params, prems, context = ctxt', ...} =>
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   363
   let
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   364
     val (prems1, prems2) = chop (length prems - length rules) prems;
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   365
     val (params1, params2) = chop (length params - length preds) params;
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   366
   in
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   367
     rtac (ObjectLogic.rulify (all_elims params1 (nth prems2 i))) 1 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   368
     THEN
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   369
     EVERY1 (map (fn prem => subproof2 prem params2 prems2 ctxt') prems1)
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   370
   end)*}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   371
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   372
ML{*
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   373
fun introductions_tac defs rules preds i ctxt =
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   374
  EVERY1 [ObjectLogic.rulify_tac,
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   375
          K (rewrite_goals_tac defs),
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   376
          REPEAT o (resolve_tac [@{thm allI},@{thm impI}]),
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   377
          subproof1 rules preds i ctxt]*}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   378
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   379
lemma evenS: 
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   380
  shows "odd m \<Longrightarrow> even (Suc m)"
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   381
apply(tactic {* 
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   382
let
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   383
  val rules = [@{prop "even (0::nat)"},
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   384
                 @{prop "\<And>n::nat. odd n \<Longrightarrow> even (Suc n)"},
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   385
                 @{prop "\<And>n::nat. even n \<Longrightarrow> odd (Suc n)"}] 
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   386
  val defs = [@{thm even_def}, @{thm odd_def}]
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   387
  val preds = [@{term "even::nat\<Rightarrow>bool"}, @{term "odd::nat\<Rightarrow>bool"}]
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   388
in
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   389
  introductions_tac defs rules preds 1 @{context}
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   390
end *})
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   391
done
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   392
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   393
ML{*fun introductions rules preds defs lthy = 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   394
let
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   395
  fun prove_intro (i, goal) =
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   396
    Goal.prove lthy [] [] goal
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   397
      (fn {context, ...} => introductions_tac defs rules preds i context)
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   398
in
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   399
  map_index prove_intro rules
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   400
end*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   401
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   402
text {* main internal function *}
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   403
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   404
ML %linenosgray{*fun add_inductive_i pred_specs rule_specs lthy =
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   405
let
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   406
  val syns = map snd pred_specs
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   407
  val pred_specs' = map fst pred_specs
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   408
  val prednames = map fst pred_specs'
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   409
  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
   410
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   411
  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
   412
  val (attrs, rules) = split_list rule_specs    
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   413
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   414
  val (defs, lthy') = definitions rules preds prednames syns tyss lthy      
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   415
  val ind_rules = inductions rules defs preds tyss lthy' 	
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   416
  val intro_rules = introductions rules preds defs lthy'
91
667a0943c40b added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   417
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   418
  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
   419
  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
   420
in
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   421
    lthy' 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   422
    |> 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
   423
        ((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
   424
    |-> (fn intross => LocalTheory.note Thm.theoremK
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   425
         ((Binding.qualify false mut_name (Binding.name "intros"), []), maps snd intross)) 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   426
    |>> snd 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   427
    ||>> (LocalTheory.notes Thm.theoremK (map (fn (((R, _), _), th) =>
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   428
         ((Binding.qualify false (Binding.name_of R) (Binding.name "induct"),
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   429
          [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
   430
           Attrib.internal (K (RuleCases.consumes 1)),
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   431
           Attrib.internal (K (Induct.induct_pred ""))]), [([th], [])]))
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   432
          (pred_specs ~~ ind_rules)) #>> maps snd) 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   433
    |> snd
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   434
end*}
91
667a0943c40b added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   435
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   436
ML{*fun read_specification' vars specs lthy =
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   437
let 
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   438
  val specs' = map (fn (a, s) => (a, [s])) specs
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   439
  val ((varst, specst), _) = 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   440
                   Specification.read_specification vars specs' lthy
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   441
  val specst' = map (apsnd the_single) specst
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   442
in   
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   443
  (varst, specst')
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   444
end*}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   445
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   446
ML{*fun add_inductive pred_specs rule_specs lthy =
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   447
let
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   448
  val (pred_specs', rule_specs') = 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   449
    read_specification' pred_specs rule_specs lthy
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   450
in
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   451
  add_inductive_i pred_specs' rule_specs' lthy
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   452
end*} 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   453
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   454
ML{*val spec_parser = 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   455
   OuterParse.opt_target --
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   456
   OuterParse.fixes -- 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   457
   Scan.optional 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   458
     (OuterParse.$$$ "where" |--
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   459
        OuterParse.!!! 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   460
          (OuterParse.enum1 "|" 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   461
             (SpecParse.opt_thm_name ":" -- OuterParse.prop))) []*}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   462
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   463
ML{*val specification =
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   464
  spec_parser >>
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   465
    (fn ((loc, pred_specs), rule_specs) =>
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   466
      Toplevel.local_theory loc (add_inductive pred_specs rule_specs))*}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   467
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   468
ML{*val _ = OuterSyntax.command "simple_inductive" "define inductive predicates"
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   469
          OuterKeyword.thy_decl specification*}
91
667a0943c40b added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   470
124
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   471
text {*
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   472
  Things to include at the end:
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   473
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   474
  \begin{itemize}
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   475
  \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
   476
  the rules
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   477
  \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
   478
  what the standard inductive package generates)
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   479
  \end{itemize}
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   480
  
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   481
*}
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   482
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   483
simple_inductive
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   484
  Even and Odd
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   485
where
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   486
  Even0: "Even 0"
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   487
| EvenS: "Odd n \<Longrightarrow> Even (Suc n)"
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   488
| 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
   489
91
667a0943c40b added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   490
end