CookBook/Package/Ind_Code.thy
author Christian Urban <urbanc@in.tum.de>
Mon, 16 Mar 2009 00:12:32 +0100
changeset 178 fb8f22dd8ad0
parent 177 4e2341f6599d
child 179 75381fa516cd
permissions -rw-r--r--
adapted to latest Attrib.setup changes and 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
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
  We use the following wrapper function to actually make the definition via
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 {*
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
    59
  Returns the definition (as theorem) and the local theory in which this definition has 
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
    60
  been made. The @{ML internalK in Thm} in Line 4 is just a flag attached to the 
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
    61
  theorem (others flags are @{ML definitionK in Thm} or @{ML axiomK in Thm}). 
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
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
    64
  use @{ML empty_binding in Attrib} in Line 3, since the definition does not need any
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
    65
  theorem attributes. 
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    66
*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    67
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
    68
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
    69
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
    70
  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
    71
  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
    72
  val _ = warning (str_of_thm lthy' def)
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
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
    74
  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
    75
end *}
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    76
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    77
text {*
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
    78
  Prints out the theorem @{prop "MyTrue \<equiv> True"}.
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
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    81
text {*
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    82
  Constructs the term for the definition: the main arguments are a predicate
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    83
  and the types of the arguments, it also expects orules which are the 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    84
  intro rules in the object logic; preds which are all predicates. returns the
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    85
  term.
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    86
*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    87
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
    88
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
    89
let 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    90
  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
    91
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    92
  val fresh_args = 
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
    93
        arg_tys 
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    94
        |> map (pair "z")
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    95
        |> Variable.variant_frees lthy orules 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    96
        |> map Free
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    97
in
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    98
  list_comb (pred, fresh_args)
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
    99
  |> fold_rev (curry HOLogic.mk_imp) orules
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   100
  |> fold_rev mk_all preds
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   101
  |> fold_rev lambda fresh_args 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   102
end*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   103
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   104
text {*
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   105
  The lines 5-9 produce fresh arguments with which the predicate can be applied.
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   106
  For this it pairs every type with a string @{text [quotes] "z"} (Line 7); then 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   107
  generates variants for all these strings (names) so that they are unique w.r.t.~to 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   108
  the intro rules; in Line 9 it generates the corresponding variable terms for these 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   109
  unique names.
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   110
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   111
  The unique free variables are applied to the predicate (Line 11); then
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   112
  the intro rules are attached as preconditions (Line 12); in Line 13 we
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   113
  quantify over all predicates; and in line 14 we just abstract over all
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   114
  the (fresh) arguments of the predicate.
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   115
*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   116
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   117
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   118
local_setup {*
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   119
fn lthy =>
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   120
let
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   121
  val orules = [@{prop "even 0"},
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   122
                @{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
   123
                @{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
   124
  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
   125
  val pred = @{term "even::nat\<Rightarrow>bool"}
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   126
  val arg_tys = [@{typ "nat"}]
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   127
  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
   128
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
   129
  warning (Syntax.string_of_term lthy def); lthy
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   130
end*}
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   131
91
667a0943c40b added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
text {*
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   133
  The arguments of the main function for the definitions are 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   134
  the intro rules; the predicates and their names, their syntax 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   135
  annotations and the argument types of each predicate. It  
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   136
  returns a theorem list (the definitions) and a local
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   137
  theory where the definitions are made
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   138
*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   139
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   140
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
   141
let
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   142
  val thy = ProofContext.theory_of lthy
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   143
  val orules = map (ObjectLogic.atomize_term thy) rules
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   144
  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
   145
in
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   146
  fold_map make_defs (prednames ~~ syns ~~ defs) lthy
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   147
end*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   148
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   149
text {*
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   150
  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
   151
  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
   152
  call @{ML defs_aux} to generate the terms for the left-hand sides.
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   153
  The actual definitions are made in Line 7.  
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
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
   156
local_setup {*
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
   157
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
   158
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
   159
  val rules = [@{prop "even 0"},
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
   160
                @{prop "\<And>n::nat. odd n \<Longrightarrow> even (Suc n)"},
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
   161
                @{prop "\<And>n::nat. even n \<Longrightarrow> odd (Suc n)"}] 
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
   162
  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
   163
  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
   164
  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
   165
  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
   166
  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
   167
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
   168
  warning (str_of_thms lthy' defs); 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
   169
end*}
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
   170
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
   171
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   172
subsection {* Induction Principles *}
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   173
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   174
text {* Recall the proof for the induction principle for @{term "even"}: *}
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   175
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   176
lemma 
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   177
  assumes prems: "even n"
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   178
  shows "P 0 \<Longrightarrow> 
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   179
             (\<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
   180
apply(atomize (full))
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   181
apply(cut_tac prems)
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   182
apply(unfold even_def)
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   183
apply(drule spec[where x=P])
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   184
apply(drule spec[where x=Q])
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   185
apply(assumption)
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   186
done
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   187
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   188
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   189
text {* We need to be able to instantiate universal quantifiers. *}
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   190
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   191
ML{*fun inst_spec ct = 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   192
 Drule.instantiate' [SOME (ctyp_of_term ct)] [NONE, SOME ct] @{thm spec}*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   193
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   194
text {*
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   195
  Instantiates the @{text "?x"} in @{thm spec} with a @{ML_type cterm}.
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   196
*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   197
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   198
text {*
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   199
  Instantiates universal qantifications in the premises
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   200
*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   201
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   202
lemma "\<forall>x\<^isub>1 x\<^isub>2 x\<^isub>3. P (x\<^isub>1::nat) (x\<^isub>2::nat) (x\<^isub>3::nat) \<Longrightarrow> True"
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   203
apply (tactic {* 
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   204
  let 
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   205
    val ctrms = [@{cterm "y\<^isub>1::nat"},@{cterm "y\<^isub>2::nat"},@{cterm "y\<^isub>3::nat"}]
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   206
  in 
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   207
    EVERY1 (map (dtac o inst_spec) ctrms)
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   208
  end *})
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   209
txt {* \begin{minipage}{\textwidth}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   210
       @{subgoals} 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   211
       \end{minipage}*}
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   212
(*<*)oops(*>*)
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   213
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   214
text {*
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   215
  Now the tactic for proving the induction rules: 
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
   216
*}
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
   217
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   218
ML{*fun induction_tac defs prems insts =
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   219
  EVERY1 [ObjectLogic.full_atomize_tac,
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   220
          cut_facts_tac prems,
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   221
          K (rewrite_goals_tac defs),
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   222
          EVERY' (map (dtac o inst_spec) insts),
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   223
          assume_tac]*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   224
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   225
lemma 
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   226
  assumes prems: "even n"
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   227
  shows "P 0 \<Longrightarrow> 
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   228
           (\<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
   229
apply(tactic {* 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   230
  let
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   231
     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
   232
     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
   233
  in 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   234
    induction_tac defs @{thms prems} insts 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   235
  end *})
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   236
done
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   237
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   238
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   239
text {*
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   240
  While the generic proof is relatively simple, it is a bit harder to
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   241
  set up the goals for the induction principles. 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   242
*}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   243
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
   244
ML {*
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
   245
fun prove_induction lthy2 defs rules cnewpreds ((pred, newpred), 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
   246
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
   247
  val zs = replicate (length tys) "z"
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
   248
  val (newargnames, lthy3) = Variable.variant_fixes zs lthy2;
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
   249
  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
   250
  
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
   251
  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
   252
  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
   253
         (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
   254
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
   255
  Goal.prove lthy3 [] [prem] goal
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
   256
  (fn {prems, ...} => induction_tac defs prems 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
   257
  |> singleton (ProofContext.export lthy3 lthy2)
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
   258
end 
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
   259
*}
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   260
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   261
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
   262
let
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   263
  val Ps = replicate (length preds) "P"
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   264
  val (newprednames, lthy2) = Variable.variant_fixes Ps lthy1
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   265
  
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   266
  val thy = ProofContext.theory_of lthy2
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   267
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   268
  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
   269
  val newpreds = map Free (newprednames ~~ tyss')
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   270
  val cnewpreds = map (cterm_of thy) newpreds
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   271
  val rules' = map (subst_free (preds ~~ newpreds)) rules
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   272
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   273
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
   274
  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
   275
        (preds ~~ newpreds ~~ tyss)
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
   276
  |> ProofContext.export lthy2 lthy1
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   277
end*}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   278
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   279
ML {*
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   280
  let 
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   281
    val rules = [@{prop "even (0::nat)"},
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   282
                 @{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
   283
                 @{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
   284
    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
   285
    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
   286
    val tyss = [[@{typ "nat"}], [@{typ "nat"}]]
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   287
  in
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   288
    inductions rules defs preds tyss @{context}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   289
  end
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   290
*}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   291
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   292
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   293
subsection {* Introduction Rules *}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   294
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   295
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
   296
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
   297
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   298
ML{*fun subproof2 prem params2 prems2 =  
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   299
 SUBPROOF (fn {prems, ...} =>
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   300
   let
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   301
     val prem' = prems MRS prem;
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   302
     val prem'' = 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   303
       case prop_of prem' of
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   304
           _ $ (Const (@{const_name All}, _) $ _) =>
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   305
             prem' |> all_elims params2 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   306
                   |> imp_elims prems2
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   307
         | _ => prem';
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   308
   in 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   309
     rtac prem'' 1 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   310
   end)*}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   311
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   312
ML{*fun subproof1 rules preds i = 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   313
 SUBPROOF (fn {params, prems, context = ctxt', ...} =>
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   314
   let
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   315
     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
   316
     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
   317
   in
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   318
     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
   319
     THEN
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   320
     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
   321
   end)*}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   322
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   323
ML{*
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   324
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
   325
  EVERY1 [ObjectLogic.rulify_tac,
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   326
          K (rewrite_goals_tac defs),
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   327
          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
   328
          subproof1 rules preds i ctxt]*}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   329
173
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   330
lemma evenS: 
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   331
  shows "odd m \<Longrightarrow> even (Suc m)"
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   332
apply(tactic {* 
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   333
let
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   334
  val rules = [@{prop "even (0::nat)"},
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   335
                 @{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
   336
                 @{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
   337
  val defs = [@{thm even_def}, @{thm odd_def}]
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   338
  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
   339
in
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   340
  introductions_tac defs rules preds 1 @{context}
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   341
end *})
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   342
done
d820cb5873ea used latex package boxedminipage
Christian Urban <urbanc@in.tum.de>
parents: 165
diff changeset
   343
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   344
ML{*fun introductions rules preds defs lthy = 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   345
let
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   346
  fun prove_intro (i, goal) =
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   347
    Goal.prove lthy [] [] goal
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   348
      (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
   349
in
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   350
  map_index prove_intro rules
164
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   351
end*}
3f617d7a2691 more work on simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 163
diff changeset
   352
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   353
text {* main internal function *}
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   354
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   355
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
   356
let
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   357
  val syns = map snd pred_specs
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   358
  val pred_specs' = map fst pred_specs
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   359
  val prednames = map fst pred_specs'
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   360
  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
   361
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   362
  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
   363
  val (attrs, rules) = split_list rule_specs    
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   364
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   365
  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
   366
  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
   367
  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
   368
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   369
  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
   370
  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
   371
in
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   372
    lthy' 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   373
    |> 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
   374
        ((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
   375
    |-> (fn intross => LocalTheory.note Thm.theoremK
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   376
         ((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
   377
    |>> snd 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   378
    ||>> (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
   379
         ((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
   380
          [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
   381
           Attrib.internal (K (RuleCases.consumes 1)),
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   382
           Attrib.internal (K (Induct.induct_pred ""))]), [([th], [])]))
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   383
          (pred_specs ~~ ind_rules)) #>> maps snd) 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   384
    |> snd
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   385
end*}
91
667a0943c40b added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   386
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   387
ML{*fun read_specification' vars specs lthy =
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   388
let 
176
3da5f3f07d8b updated to new read_specification
Christian Urban <urbanc@in.tum.de>
parents: 173
diff changeset
   389
  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
   390
  val ((varst, specst), _) = 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   391
                   Specification.read_specification vars specs' lthy
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   392
  val specst' = map (apsnd the_single) specst
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   393
in   
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   394
  (varst, specst')
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   395
end*}
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   396
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   397
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
   398
let
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   399
  val (pred_specs', rule_specs') = 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   400
    read_specification' pred_specs rule_specs lthy
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   401
in
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   402
  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
   403
end*} 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   404
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   405
ML{*val spec_parser = 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   406
   OuterParse.opt_target --
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   407
   OuterParse.fixes -- 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   408
   Scan.optional 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   409
     (OuterParse.$$$ "where" |--
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   410
        OuterParse.!!! 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   411
          (OuterParse.enum1 "|" 
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   412
             (SpecParse.opt_thm_name ":" -- OuterParse.prop))) []*}
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
ML{*val specification =
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   415
  spec_parser >>
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   416
    (fn ((loc, pred_specs), rule_specs) =>
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   417
      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
   418
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   419
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
   420
          OuterKeyword.thy_decl specification*}
91
667a0943c40b added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   421
124
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   422
text {*
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   423
  Things to include at the end:
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   424
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   425
  \begin{itemize}
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   426
  \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
   427
  the rules
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   428
  \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
   429
  what the standard inductive package generates)
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   430
  \end{itemize}
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   431
  
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   432
*}
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 118
diff changeset
   433
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   434
simple_inductive
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   435
  Even and Odd
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   436
where
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   437
  Even0: "Even 0"
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   438
| EvenS: "Odd n \<Longrightarrow> Even (Suc n)"
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 164
diff changeset
   439
| 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
   440
91
667a0943c40b added a section that will eventually describe the code
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   441
end