ProgTutorial/Package/Ind_General_Scheme.thy
author Norbert Schirmer <norbert.schirmer@web.de>
Fri, 17 May 2019 11:21:09 +0200
changeset 571 95b42288294e
parent 565 cecd7a941885
permissions -rw-r--r--
reactivated Readme.thy for authors
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 
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 295
diff changeset
     2
imports  Ind_Intro Simple_Inductive_Package
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
244
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
     5
(*<*)
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
     6
simple_inductive
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
     7
  trcl :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
     8
where
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
     9
  base: "trcl R x x"
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    10
| step: "trcl R x y \<Longrightarrow> R y z \<Longrightarrow> trcl R x z"
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    11
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    12
simple_inductive
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    13
  even and odd
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    14
where
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    15
  even0: "even 0"
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    16
| evenS: "odd n \<Longrightarrow> even (Suc n)"
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    17
| oddS: "even n \<Longrightarrow> odd (Suc n)"
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    18
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    19
simple_inductive
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    20
  accpart :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool"
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    21
where
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    22
  accpartI: "(\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x"
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    23
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    24
datatype trm =
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    25
  Var "string"
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    26
| App "trm" "trm"
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    27
| Lam "string" "trm"
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    28
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    29
simple_inductive 
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    30
  fresh :: "string \<Rightarrow> trm \<Rightarrow> bool" 
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    31
where
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    32
  fresh_var: "a\<noteq>b \<Longrightarrow> fresh a (Var b)"
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    33
| fresh_app: "\<lbrakk>fresh a t; fresh a s\<rbrakk> \<Longrightarrow> fresh a (App t s)"
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    34
| fresh_lam1: "fresh a (Lam a t)"
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    35
| fresh_lam2: "\<lbrakk>a\<noteq>b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)"
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    36
(*>*)
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    37
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 237
diff changeset
    38
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    39
section \<open>The Code in a Nutshell\label{sec:nutshell}\<close>
210
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
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    42
text \<open>
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
    43
  The inductive package will generate the reasoning infrastructure for
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    44
  mutually recursive predicates, say \<open>pred\<^sub>1\<dots>pred\<^sub>n\<close>. In what
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
    45
  follows we will have the convention that various, possibly empty collections
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 278
diff changeset
    46
  of ``things'' (lists, terms, nested implications and so on) are indicated either by
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
    47
  adding an @{text [quotes] "s"} or by adding a superscript @{text [quotes]
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    48
  "\<^sup>*"}. The shorthand for the predicates will therefore be \<open>preds\<close> or \<open>pred\<^sup>*\<close>. In the case of the predicates there must
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
    49
  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
    50
  definition.
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    51
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    52
  The input for the inductive package will be some \<open>preds\<close> with possible 
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    53
  typing and syntax annotations, and also some introduction rules. We call below the 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    54
  introduction rules short as \<open>rules\<close>. Borrowing some idealised Isabelle 
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    55
  notation, one such \<open>rule\<close> 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
    56
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    57
  \begin{isabelle}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    58
  \<open>rule ::= 
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    59
  \<And>xs. \<^latex>\<open>$\underbrace{\mbox{\<close>As\<^latex>\<open>}}_{\text{\makebox[0mm]{\rm non-recursive premises}}}$\<close> \<Longrightarrow> 
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    60
  \<^latex>\<open>$\underbrace{\mbox{\<close>(\<And>ys. Bs \<Longrightarrow> pred ss)\<^sup>*\<^latex>\<open>}}_{\text{\rm recursive premises}}$>\<close> 
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    61
  \<Longrightarrow> pred ts\<close>
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    62
  \end{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    63
  
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    64
  For the purposes here, we will assume the \<open>rules\<close> have this format and
278
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
    65
  omit any code that actually tests this. Therefore ``things'' can go horribly
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    66
  wrong, if the \<open>rules\<close> are not of this form.  The \<open>As\<close> and
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    67
  \<open>Bs\<close> in a \<open>rule\<close> stand for formulae not involving the
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    68
  inductive predicates \<open>preds\<close>; the instances \<open>pred ss\<close> and
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    69
  \<open>pred ts\<close> can stand for different predicates, like \<open>pred\<^sub>1 ss\<close> and \<open>pred\<^sub>2 ts\<close>, in case mutual recursive
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    70
  predicates are defined; the terms \<open>ss\<close> and \<open>ts\<close> are the
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 278
diff changeset
    71
  arguments of these predicates. Every formula left of @{text [quotes] "\<Longrightarrow> pred
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    72
  ts"} is a premise of the rule. The outermost quantified variables \<open>xs\<close> are usually omitted in the user's input. The quantification for the
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    73
  variables \<open>ys\<close> is local with respect to one recursive premise and
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    74
  must be given. Some examples of \<open>rule\<close>s are
278
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
    75
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    76
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    77
  @{thm [display] fresh_var[no_vars]}
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
  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
    80
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    81
  @{thm [display] evenS[no_vars]}
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
  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
    84
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    85
  @{thm [display] accpartI[no_vars]}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    86
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 278
diff changeset
    87
  has a single recursive premise that has a precondition. As is custom all 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    88
  rules are stated without the leading meta-quantification \<open>\<And>xs\<close>.
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    89
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
    90
  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
    91
  induction principles and introduction rules.  For the definitions we need to have the
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    92
  \<open>rules\<close> in a form where the meta-quantifiers and meta-implications are
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    93
  replaced by their object logic equivalents. Therefore an \<open>orule\<close> is
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
    94
  of the form
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    95
551
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 517
diff changeset
    96
  @{text [display] "orule ::= \<forall>xs. As \<longrightarrow> (\<forall>ys. Bs \<longrightarrow> pred ss)\<^sup>* \<longrightarrow> pred ts"}
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    97
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    98
  A definition for the predicate \<open>pred\<close> has then the form
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    99
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   100
  @{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
   101
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   102
  The induction principles for every predicate \<open>pred\<close> are of the
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   103
  form
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
  @{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
   106
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   107
  where in the \<open>rules\<close>-part every \<open>pred\<close> is replaced by a fresh
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   108
  schematic variable \<open>?P\<close>.
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   109
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   110
  In order to derive an induction principle for the predicate \<open>pred\<close>,
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   111
  we first transform \<open>ind\<close> into the object logic and fix the schematic variables. 
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
   112
  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
   113
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   114
  @{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
   115
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   116
  If we assume \<open>pred zs\<close> and unfold its definition, then we have an
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   117
  assumption
210
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] "\<forall>preds. orules \<longrightarrow> pred zs"} 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   120
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   121
  and must prove the goal
210
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
  @{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
   124
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   125
  This can be done by instantiating the \<open>\<forall>preds\<close>-quantification 
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   126
  with the \<open>Ps\<close>. Then we are done since we are left with a simple 
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
   127
  identity.
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   128
  
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   129
  Although the user declares the introduction rules \<open>rules\<close>, they must 
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   130
  also be derived from the \<open>defs\<close>. 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
   131
  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
   132
551
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 517
diff changeset
   133
  @{text [display] "\<And>xs. As \<Longrightarrow> (\<And>ys. Bs \<Longrightarrow> pred ss)\<^sup>* \<Longrightarrow> pred ts"}
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   134
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   135
  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
   136
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   137
  \begin{isabelle}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   138
  (i)~~\<open>As\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   139
  (ii)~\<open>(\<And>ys. Bs \<Longrightarrow> pred ss)\<^sup>*\<close>
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   140
  \end{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   141
212
ac01ddb285f6 polishing
Christian Urban <urbanc@in.tum.de>
parents: 211
diff changeset
   142
  and must show the goal
210
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
  @{text [display] "pred ts"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   145
  
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   146
  If we now unfold the definitions for the \<open>preds\<close>, we have assumptions
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   147
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   148
  \begin{isabelle}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   149
  (i)~~~\<open>As\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   150
  (ii)~~\<open>(\<And>ys. Bs \<Longrightarrow> \<forall>preds. orules \<longrightarrow> pred ss)\<^sup>*\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   151
  (iii)~\<open>orules\<close>
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   152
  \end{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   153
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   154
  and need to show
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   155
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   156
  @{text [display] "pred ts"}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   157
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   158
  In the last step we removed some quantifiers and moved the precondition \<open>orules\<close>  
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   159
  into the assumption. The \<open>orules\<close> stand for all introduction rules that are given 
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   160
  by the user. We apply the \<open>orule\<close> that corresponds to introduction rule we are 
278
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
   161
  proving. After transforming the object connectives into meta-connectives, this introduction 
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
   162
  rule must necessarily be of the form 
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   163
551
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 517
diff changeset
   164
  @{text [display] "As \<Longrightarrow> (\<And>ys. Bs \<Longrightarrow> pred ss)\<^sup>* \<Longrightarrow> pred ts"}
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   165
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   166
  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
   167
  goals of the form
210
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
  \begin{isabelle}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   170
  (a)~\<open>As\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   171
  (b)~\<open>(\<And>ys. Bs \<Longrightarrow> pred ss)\<^sup>*\<close>
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   172
  \end{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   173
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   174
  We can immediately discharge the goals \<open>As\<close> using the assumptions in 
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   175
  \<open>(i)\<close>. The goals in \<open>(b)\<close> can be discharged as follows: we 
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   176
  assume the \<open>Bs\<close> and prove \<open>pred ss\<close>. For this we resolve the 
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   177
  \<open>Bs\<close>  with the assumptions in \<open>(ii)\<close>. This gives us 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: 210
diff changeset
   178
  assumptions
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   179
551
be361e980acf updated subscripts
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 517
diff changeset
   180
  @{text [display] "(\<forall>preds. orules \<longrightarrow> pred ss)\<^sup>*"}
210
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
  Instantiating the universal quantifiers and then resolving with the assumptions 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   183
  in \<open>(iii)\<close> gives us \<open>pred ss\<close>, 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
   184
  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
   185
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   186
  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
   187
  We will describe the code in the next section. For building testcases, we use the shorthands for 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   188
  \<open>even/odd\<close>, @{term "fresh"} and @{term "accpart"}
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 212
diff changeset
   189
  defined in Figure~\ref{fig:shorthands}.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   190
\<close>
210
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
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   193
text_raw\<open>
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   194
\begin{figure}[p]
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   195
\begin{minipage}{\textwidth}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   196
\begin{isabelle}\<close>  
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   197
ML %grayML\<open>(* even-odd example *)
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   198
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
   199
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   200
val eo_rules =  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   201
  [@{prop "even 0"},
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   202
   @{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
   203
   @{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
   204
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   205
val eo_orules =  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   206
  [@{prop "even 0"},
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   207
   @{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
   208
   @{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
   209
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 278
diff changeset
   210
val eo_preds =  [@{term "even::nat \<Rightarrow> bool"}, @{term "odd::nat \<Rightarrow> bool"}] 
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   211
val eo_prednames = [@{binding "even"}, @{binding "odd"}]
237
0a8981f52045 very slight polishing to the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 219
diff changeset
   212
val eo_mxs = [NoSyn, NoSyn] 
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   213
val eo_arg_tyss = [[@{typ "nat"}], [@{typ "nat"}]] 
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 278
diff changeset
   214
val e_pred = @{term "even::nat \<Rightarrow> bool"}
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   215
val e_arg_tys = [@{typ "nat"}] 
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
(* freshness example *)
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   220
val fresh_rules =  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   221
  [@{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
   222
   @{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
   223
   @{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
   224
   @{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
   225
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   226
val fresh_orules =  
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   227
  [@{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
   228
   @{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
   229
   @{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
   230
   @{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
   231
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 278
diff changeset
   232
val fresh_pred =  @{term "fresh::string \<Rightarrow> trm \<Rightarrow> bool"} 
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   233
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
   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
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   237
(* accessible-part example *)
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   238
val acc_rules = 
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   239
     [@{prop "\<And>R x. (\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x"}]
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   240
val acc_pred = @{term "accpart::('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow>'a \<Rightarrow> bool"}\<close>
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   241
text_raw\<open>
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   242
\end{isabelle}
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   243
\end{minipage}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   244
\caption{Shorthands for the inductive predicates \<open>even\<close>/\<open>odd\<close>, 
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   245
  \<open>fresh\<close> and \<open>accpart\<close>. The names of these shorthands follow 
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   246
  the convention \<open>rules\<close>, \<open>orules\<close>, \<open>preds\<close> and so on. 
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
   247
  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
   248
  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
   249
\end{figure}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   250
\<close>
210
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   251
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   252
db8e302f44c8 more work on the simple inductive section
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   253
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   254
end