ProgTutorial/Package/Ind_General_Scheme.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 26 Mar 2009 19:00:51 +0000
changeset 210 db8e302f44c8
parent 189 069d525f8f1d
child 211 d5accbc67e1b
permissions -rw-r--r--
more work on the simple inductive section
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
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    20
section {* The Code in a 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 {*
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    23
  (FIXME: perhaps move somewhere else)
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    24
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
    25
  The point of these examples is to get a feeling what the automatic proofs 
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
    26
  should do in order to solve all inductive definitions we throw at them. For this 
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
    27
  it is instructive to look at the general construction principle 
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
    28
  of inductive definitions, which we shall do in the next section.
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    29
*}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    30
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    31
text {*
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    32
  The inductive package will generate the reasoning infrastructure
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    33
  for mutually recursive predicates @{text "pred\<^isub>1\<dots>pred\<^isub>n"}. In what
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    34
  follows we will have the convention that various, possibly empty collections of 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    35
  ``things'' are indicated either by adding an @{text "s"} or by adding 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    36
  a superscript @{text [quotes] "\<^isup>*"}. The shorthand for the predicates 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    37
  will therefore be @{text "preds"} or @{text "pred\<^sup>*"}. In this case of the
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    38
  predicates there must be at least a single one in order to obtain a meaningful
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    39
  definition.
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    40
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    41
  The input for the inductive predicate will be some @{text "preds"} with possible 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    42
  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
    43
  introduction rules short as @{text "rules"}. Borrowing some idealised Isabelle 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    44
  notation, one such @{text "rule"} is of the form
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    45
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    46
  \begin{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    47
  @{text "rule ::= 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    48
  \<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
    49
  \<^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
    50
  \<Longrightarrow> pred ts"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    51
  \end{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    52
  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    53
  We actually assume the user will always state rules of this form and we omit
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    54
  any code that test this format. So things can go horribly wrong if the
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    55
  @{text "rules"} are not of this form.\footnote{FIXME: Exercise to test this
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    56
  format.} The @{text As} and @{text Bs} in a @{text "rule"} are formulae not
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    57
  involving the inductive predicates @{text "preds"}; the instances @{text
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    58
  "pred ss"} and @{text "pred ts"} can stand for different predicates.
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    59
  Everything left of @{text [quotes] "\<Longrightarrow> pred ts"} are called the premises of
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    60
  the rule. The variables @{text "xs"} are usually omitted in the user's
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    61
  input. The variables @{text "ys"} are local with respect to on recursive
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    62
  premise. Some examples of @{text "rule"}s the user might write are:
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
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    65
  @{thm [display] fresh_var[no_vars]}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    66
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    67
  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
    68
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    69
  @{thm [display] evenS[no_vars]}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    70
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    71
  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
    72
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    73
  @{thm [display] accpartI[no_vars]}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    74
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    75
  has a recursive premise which has a precondition. In the examples, all 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    76
  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
    77
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    78
  The code of the inductive package falls roughly in tree parts involving
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    79
  the definitions, the induction principles and the introduction rules, 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    80
  respectively. For the definitions we need to have the @{text rules}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    81
  in a form where the meta-quantifiers and meta-implications are replaced
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    82
  by their object logic equivalent. Therefore an @{text "orule"} is of the
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    83
  form 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    84
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    85
  @{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
    86
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    87
  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
    88
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    89
  @{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
    90
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    91
  The induction principles for the predicate @{text "pred"} are of the
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    92
  form
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    93
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    94
  @{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
    95
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    96
  where in the @{text "rules"} every @{text pred} is replaced by a new
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    97
  (meta)variable @{text "?P"}.
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    98
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    99
  In order to derive an induction principle for the predicate @{text "pred"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   100
  we first transform it into the object logic and fix the meta-variables. Hence 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   101
  we have to prove a formula of the form
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   102
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   103
  @{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
   104
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   105
  If we assume @{text "pred zs"} and unfold its definition, then we have
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   106
  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   107
  @{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
   108
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   109
  and must prove
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   110
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   111
  @{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
   112
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   113
  This can be done by instantiating the @{text "\<forall>preds"} with the @{text "Ps"}. 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   114
  Then we are done since we are left with a simple identity.
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
  The proofs for the introduction rules are more involved. Assuming we want to
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   117
  prove the introduction rule 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   118
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   119
  @{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
   120
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   121
  then we can assume
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   122
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   123
  \begin{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   124
  (i)~~@{text "As"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   125
  (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
   126
  \end{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   127
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   128
  and must show
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   129
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   130
  @{text [display] "pred ts"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   131
  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   132
  If we now unfold the definitions for the @{text preds}, we have
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   133
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   134
  \begin{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   135
  (i)~~~@{text "As"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   136
  (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
   137
  (iii)~@{text "orules"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   138
  \end{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   139
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   140
  and need to show
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   141
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   142
  @{text [display] "pred ts"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   143
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   144
  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
   145
  into the assumtion. The @{text "orules"} stand for all introduction rules that are given 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   146
  by the user. We apply the one which corresponds to introduction rule we are proving.
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   147
  This introduction rule must be of the form 
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
  @{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
   150
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   151
  When we apply this rule we end up in the goal state where we have to prove
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   152
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   153
  \begin{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   154
  (a)~@{text "As"}\\
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   155
  (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
   156
  \end{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   157
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   158
  The goals @{text "As"} we can immediately discharge with the assumption in @{text "(i)"}.
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   159
  The goals in @{text "(b)"} we discharge as follows: we assume the @{text "(iv)"} 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   160
  @{text "Bs"} and prove @{text "(c)"} @{text "pred ss"}. We then resolve the 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   161
  @{text "Bs"}  with the assumptions in @{text "(ii)"}. This gives us
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.
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   167
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   168
  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   169
  What remains is to implement the reasoning outlined above. 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   170
  For building testcases, we use some shorthands for the definitions 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   171
  of @{text "even/odd"}, @{term "fresh"} and @{term "accpart"}. 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   172
*}
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
text_raw{*
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   176
\begin{figure}[p]
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   177
\begin{minipage}{\textwidth}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   178
\begin{isabelle}*}  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   179
ML{*(* even-odd example *)
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   180
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
   181
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   182
val eo_rules =  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   183
  [@{prop "even 0"},
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   184
   @{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
   185
   @{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
   186
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   187
val eo_orules =  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   188
  [@{prop "even 0"},
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   189
   @{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
   190
   @{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
   191
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   192
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
   193
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
   194
val eo_syns = [NoSyn, NoSyn] 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   195
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
   196
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
   197
val e_arg_tys = [@{typ "nat"}] 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   198
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
(* freshness example *)
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   202
val fresh_rules =  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   203
  [@{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
   204
   @{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
   205
   @{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
   206
   @{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
   207
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   208
val fresh_orules =  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   209
  [@{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
   210
   @{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
   211
   @{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
   212
   @{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
   213
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   214
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
   215
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
   216
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
(* accessible-part example *)
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   220
val acc_rules = 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   221
     [@{prop "\<And>R x. (\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x"}]
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   222
val acc_pred = @{term "accpart:: ('a \<Rightarrow>'a\<Rightarrow>bool)\<Rightarrow>'a \<Rightarrow>bool"}*}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   223
text_raw{*
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   224
\end{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   225
\end{minipage}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   226
\caption{Shorthands for the inductive predicates of @{text "even"}-@{text "odd"}, 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   227
  @{text "fresh"} and @{text "accpart"}. The follow the convention @{text "rules"}, 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   228
  @{text "orules"}, @{text "preds"} and so on as used in Section ???. The purpose 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   229
  of these shorthands is to simplify the construction of testcases.}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   230
\end{figure}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   231
*}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   232
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
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   235
end