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