ProgTutorial/Package/Ind_General_Scheme.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 31 Mar 2009 16:50:13 +0100
changeset 217 75154f4d4e2f
parent 215 8d1a344a621e
child 219 98d43270024f
permissions -rw-r--r--
used antiquotations
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
     1
theory Ind_General_Scheme 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
     2
imports Simple_Inductive_Package Ind_Prelims
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     3
begin
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     4
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
     5
(*<*)
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
     6
datatype trm =
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
     7
  Var "string"
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
     8
| App "trm" "trm"
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
     9
| Lam "string" "trm"
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    10
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    11
simple_inductive 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    12
  fresh :: "string \<Rightarrow> trm \<Rightarrow> bool" 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    13
where
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    14
  fresh_var: "a\<noteq>b \<Longrightarrow> fresh a (Var b)"
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    15
| fresh_app: "\<lbrakk>fresh a t; fresh a s\<rbrakk> \<Longrightarrow> fresh a (App t s)"
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    16
| fresh_lam1: "fresh a (Lam a t)"
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    17
| fresh_lam2: "\<lbrakk>a\<noteq>b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)"
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    18
(*>*)
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    19
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    20
section {* The Code in a Nutshell\label{sec:nutshell} *}
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    21
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    22
text {*
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
    23
  The inductive package will generate the reasoning infrastructure for
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
    24
  mutually recursive predicates @{text "pred\<^isub>1\<dots>pred\<^isub>n"}. In what
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
    25
  follows we will have the convention that various, possibly empty collections
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
    26
  of ``things'' (lists, nested implications and so on) are indicated either by
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
    27
  adding an @{text [quotes] "s"} or by adding a superscript @{text [quotes]
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
    28
  "\<^isup>*"}. The shorthand for the predicates will therefore be @{text
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
    29
  "preds"} or @{text "pred\<^sup>*"}. In the case of the predicates there must
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
    30
  be, of course, at least a single one in order to obtain a meaningful
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
    31
  definition.
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    32
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    33
  The input for the inductive package will be some @{text "preds"} with possible 
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    34
  typing and syntax annotations, and also some introduction rules. We call below the 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    35
  introduction rules short as @{text "rules"}. Borrowing some idealised Isabelle 
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    36
  notation, one such @{text "rule"} is assumed to be of the form
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    37
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    38
  \begin{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    39
  @{text "rule ::= 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    40
  \<And>xs. \<^raw:$\underbrace{\mbox{>As\<^raw:}}_{\text{\makebox[0mm]{\rm non-recursive premises}}}$> \<Longrightarrow> 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    41
  \<^raw:$\underbrace{\mbox{>(\<And>ys. Bs \<Longrightarrow> pred ss)\<^isup>*\<^raw:}}_{\text{\rm recursive premises}}$> 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    42
  \<Longrightarrow> pred ts"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    43
  \end{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    44
  
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    45
  For the purposes here, we will assume the @{text rules} have
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    46
  this format and omit any code that actually tests this. Therefore ``things'' 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    47
  can go horribly wrong, if the @{text "rules"} are not of this
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    48
  form.\footnote{FIXME: Exercise to test this format.} The @{text As} and
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
    49
  @{text Bs} in a @{text "rule"} stand for formulae not involving the inductive
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    50
  predicates @{text "preds"}; the instances @{text "pred ss"} and @{text "pred
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    51
  ts"} can stand for different predicates, like @{text "pred\<^isub>1 ss"} and 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    52
  @{text "pred\<^isub>2 ts"}; @{text ss} and @{text ts} are the
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
    53
  arguments of these predicates. Every formula left of @{text [quotes]
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    54
  "\<Longrightarrow> pred ts"} is a premise of the rule. The outermost quantified 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    55
  variables @{text "xs"} are usually omitted in the user's input. The 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    56
  quantification for the variables @{text "ys"} is local with respect to 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    57
  one recursive premise and must be given. Some examples of @{text "rule"}s 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    58
  are
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    59
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    60
  @{thm [display] fresh_var[no_vars]}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    61
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    62
  which has only a single non-recursive premise, whereas
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    63
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    64
  @{thm [display] evenS[no_vars]}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    65
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    66
  has a single recursive premise; the rule
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    67
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    68
  @{thm [display] accpartI[no_vars]}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    69
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
    70
  has a single recursive premise that has a precondition. As usual all 
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    71
  rules are stated without the leading meta-quantification @{text "\<And>xs"}.
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    72
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
    73
  The output of the inductive package will be definitions for the predicates, 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
    74
  induction principles and introduction rules.  For the definitions we need to have the
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
    75
  @{text rules} in a form where the meta-quantifiers and meta-implications are
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
    76
  replaced by their object logic equivalents. Therefore an @{text "orule"} is
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
    77
  of the form
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    78
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    79
  @{text [display] "orule ::= \<forall>xs. As \<longrightarrow> (\<forall>ys. Bs \<longrightarrow> pred ss)\<^isup>* \<longrightarrow> pred ts"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    80
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    81
  A definition for the predicate @{text "pred"} has then the form
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    82
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    83
  @{text [display] "def ::= pred \<equiv> \<lambda>zs. \<forall>preds. orules \<longrightarrow> pred zs"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    84
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    85
  The induction principles for every predicate @{text "pred"} are of the
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    86
  form
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    87
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    88
  @{text [display] "ind ::= pred ?zs \<Longrightarrow> rules[preds := ?Ps] \<Longrightarrow> ?P ?zs"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    89
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    90
  where in the @{text "rules"} every @{text pred} is replaced by a fresh
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    91
  meta-variable @{text "?P"}.
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    92
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    93
  In order to derive an induction principle for the predicate @{text "pred"},
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    94
  we first transform @{text ind} into the object logic and fix the meta-variables. 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
    95
  Hence we have to prove a formula of the form
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    96
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    97
  @{text [display] "pred zs \<longrightarrow> orules[preds := Ps] \<longrightarrow> P zs"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    98
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
    99
  If we assume @{text "pred zs"} and unfold its definition, then we have an
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   100
  assumption
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   101
  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   102
  @{text [display] "\<forall>preds. orules \<longrightarrow> pred zs"} 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   103
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   104
  and must prove the goal
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   105
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   106
  @{text [display] "orules[preds := Ps] \<longrightarrow> P zs"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   107
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   108
  This can be done by instantiating the @{text "\<forall>preds"}-quantification 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   109
  with the @{text "Ps"}. Then we are done since we are left with a simple 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   110
  identity.
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   111
  
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   112
  Although the user declares the introduction rules @{text rules}, they must 
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   113
  also be derived from the @{text defs}. These derivations are a bit involved. 
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   114
  Assuming we want to prove the introduction rule 
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   115
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   116
  @{text [display] "\<And>xs. As \<Longrightarrow> (\<And>ys. Bs \<Longrightarrow> pred ss)\<^isup>* \<Longrightarrow> pred ts"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   117
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   118
  then we have assumptions of the form
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   119
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   120
  \begin{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   121
  (i)~~@{text "As"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   122
  (ii)~@{text "(\<And>ys. Bs \<Longrightarrow> pred ss)\<^isup>*"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   123
  \end{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   124
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   125
  and must show the goal
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   126
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   127
  @{text [display] "pred ts"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   128
  
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   129
  If we now unfold the definitions for the @{text preds}, we have assumptions
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   130
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   131
  \begin{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   132
  (i)~~~@{text "As"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   133
  (ii)~~@{text "(\<And>ys. Bs \<Longrightarrow> \<forall>preds. orules \<longrightarrow> pred ss)\<^isup>*"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   134
  (iii)~@{text "orules"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   135
  \end{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   136
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   137
  and need to show
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   138
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   139
  @{text [display] "pred ts"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   140
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   141
  In the last step we removed some quantifiers and moved the precondition @{text "orules"}  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   142
  into the assumtion. The @{text "orules"} stand for all introduction rules that are given 
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   143
  by the user. We apply the @{text orule} that corresponds to introduction rule we are 
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   144
  proving. After lifting to the meta-connectives, this introduction rule must necessarily 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   145
  be of the form 
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   146
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   147
  @{text [display] "As \<Longrightarrow> (\<And>ys. Bs \<Longrightarrow> pred ss)\<^isup>* \<Longrightarrow> pred ts"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   148
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   149
  When we apply this rule we end up in the goal state where we have to prove
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   150
  goals of the form
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   151
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   152
  \begin{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   153
  (a)~@{text "As"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   154
  (b)~@{text "(\<And>ys. Bs \<Longrightarrow> pred ss)\<^isup>*"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   155
  \end{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   156
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   157
  We can immediately discharge the goals @{text "As"} using the assumptions in 
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   158
  @{text "(i)"}. The goals in @{text "(b)"} can be discharged as follows: we 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   159
  assume the @{text "Bs"} and prove @{text "pred ss"}. For this we resolve the 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   160
  @{text "Bs"}  with the assumptions in @{text "(ii)"}. This gives us the 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   161
  assumptions
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   162
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   163
  @{text [display] "(\<forall>preds. orules \<longrightarrow> pred ss)\<^isup>*"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   164
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   165
  Instantiating the universal quantifiers and then resolving with the assumptions 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   166
  in @{text "(iii)"} gives us @{text "pred ss"}, which is the goal we are after.
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   167
  This completes the proof for introduction rules.
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   168
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   169
  What remains is to implement in Isabelle the reasoning outlined in this section. 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   170
  We will describe the code in the next section. For building testcases, we use the shorthands for 
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   171
  @{text "even/odd"}, @{term "fresh"} and @{term "accpart"}
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   172
  defined in Figure~\ref{fig:shorthands}.
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   173
*}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   174
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   175
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   176
text_raw{*
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   177
\begin{figure}[p]
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   178
\begin{minipage}{\textwidth}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   179
\begin{isabelle}*}  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   180
ML{*(* even-odd example *)
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   181
val eo_defs = [@{thm even_def}, @{thm odd_def}]
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   182
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   183
val eo_rules =  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   184
  [@{prop "even 0"},
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   185
   @{prop "\<And>n. odd n \<Longrightarrow> even (Suc n)"},
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   186
   @{prop "\<And>n. even n \<Longrightarrow> odd (Suc n)"}]
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   187
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   188
val eo_orules =  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   189
  [@{prop "even 0"},
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   190
   @{prop "\<forall>n. odd n \<longrightarrow> even (Suc n)"},
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   191
   @{prop "\<forall>n. even n \<longrightarrow> odd (Suc n)"}]
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   192
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   193
val eo_preds =  [@{term "even::nat\<Rightarrow>bool"}, @{term "odd::nat\<Rightarrow>bool"}] 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   194
val eo_prednames = [@{binding "even"}, @{binding "odd"}]
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   195
val eo_syns = [NoSyn, NoSyn] 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   196
val eo_arg_tyss = [[@{typ "nat"}], [@{typ "nat"}]] 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   197
val e_pred = @{term "even::nat\<Rightarrow>bool"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   198
val e_arg_tys = [@{typ "nat"}] 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   199
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   200
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   201
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   202
(* freshness example *)
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   203
val fresh_rules =  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   204
  [@{prop "\<And>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: 189
diff changeset
   205
   @{prop "\<And>a s t. 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: 189
diff changeset
   206
   @{prop "\<And>a t. fresh a (Lam a t)"},
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   207
   @{prop "\<And>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: 189
diff changeset
   208
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   209
val fresh_orules =  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   210
  [@{prop "\<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: 189
diff changeset
   211
   @{prop "\<forall>a s t. 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: 189
diff changeset
   212
   @{prop "\<forall>a t. fresh a (Lam a t)"},
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   213
   @{prop "\<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: 189
diff changeset
   214
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   215
val fresh_pred =  @{term "fresh::string\<Rightarrow>trm\<Rightarrow>bool"} 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   216
val fresh_arg_tys = [@{typ "string"}, @{typ "trm"}]
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   217
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   218
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   219
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   220
(* accessible-part example *)
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   221
val acc_rules = 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   222
     [@{prop "\<And>R x. (\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x"}]
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   223
val acc_pred = @{term "accpart::('a \<Rightarrow>'a\<Rightarrow>bool)\<Rightarrow>'a \<Rightarrow>bool"}*}
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   224
text_raw{*
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   225
\end{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   226
\end{minipage}
211
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   227
\caption{Shorthands for the inductive predicates @{text "even"}-@{text "odd"}, 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   228
  @{text "fresh"} and @{text "accpart"}. The names of these shorthands follow 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   229
  the convention @{text "rules"}, @{text "orules"}, @{text "preds"} and so on. 
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   230
  The purpose of these shorthands is to simplify the construction of testcases
d5accbc67e1b more work on simple inductive and marked all sections that are still seriously incomplete with TBD
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   231
  in Section~\ref{sec:code}.\label{fig:shorthands}}
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   232
\end{figure}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   233
*}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   234
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   235
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   236
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   237
end