ProgTutorial/Package/Ind_Prelims.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 19 Nov 2009 20:23:15 +0100
changeset 397 6b423b39cc11
parent 346 0fea8b7a14a1
child 513 f223f8223d4a
permissions -rw-r--r--
tuned
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
     1
theory Ind_Prelims
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 
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
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
     5
section{* Preliminaries *}
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
     6
  
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
     7
text {*
218
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
     8
  The user will just give a specification of inductive predicate(s) and
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
     9
  expects from the package to produce a convenient reasoning
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
    10
  infrastructure. This infrastructure needs to be derived from the definition
218
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    11
  that correspond to the specified predicate(s). Before we start with
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
    12
  explaining all parts of the package, let us first give some examples showing
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
    13
  how to define inductive predicates and then also how to generate a reasoning
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
    14
  infrastructure for them. From the examples we will figure out a general
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
    15
  method for defining inductive predicates. This is usually the first step in
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
    16
  writing a package for Isabelle. The aim in this section is \emph{not} to
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    17
  write proofs that are as beautiful as possible, but as close as possible to
244
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
    18
  the ML-implementation we will develop in later sections.
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    19
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
    20
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    21
  We first consider the transitive closure of a relation @{text R}. The 
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    22
  ``pencil-and-paper'' specification for the transitive closure is:
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    23
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
    24
  \begin{center}\small
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 115
diff changeset
    25
  @{prop[mode=Axiom] "trcl R x x"} \hspace{5mm}
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 115
diff changeset
    26
  @{prop[mode=Rule] "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"}
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    27
  \end{center}
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    28
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
    29
  As mentioned before, the package has to make an appropriate definition for
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
    30
  @{term "trcl"}. Since an inductively defined predicate is the least
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
    31
  predicate closed under a collection of introduction rules, the predicate
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
    32
  @{text "trcl R x y"} can be defined so that it holds if and only if @{text
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
    33
  "P x y"} holds for every predicate @{text P} closed under the rules
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
    34
  above. This gives rise to the definition
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    35
*}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    36
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
    37
definition "trcl \<equiv> 
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
    38
     \<lambda>R x y. \<forall>P. (\<forall>x. P x x) 
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
    39
                  \<longrightarrow> (\<forall>x y z. R x y \<longrightarrow> P y z \<longrightarrow> P x z) \<longrightarrow> P x y"
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    40
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    41
text {*
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
    42
  Note we have to use the object implication @{text "\<longrightarrow>"} and object
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
    43
  quantification @{text "\<forall>"} for stating this definition (there is no other
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
    44
  way for definitions in HOL). However, the introduction rules and induction
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
    45
  principles associated with the transitive closure should use the
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
    46
  meta-connectives, since they simplify the reasoning for the user.
218
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    47
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    48
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
    49
  With this definition, the proof of the induction principle for @{term trcl}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
    50
  is almost immediate. It suffices to convert all the meta-level
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 115
diff changeset
    51
  connectives in the lemma to object-level connectives using the
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    52
  proof method @{text atomize} (Line 4 below), expand the definition of @{term trcl}
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
    53
  (Line 5 and 6), eliminate the universal quantifier contained in it (Line~7),
218
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    54
  and then solve the goal by @{text assumption} (Line 8).
113
9b6c9d172378 slightly updated
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    55
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    56
*}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    57
114
13fd0a83d3c3 properly handled linenumbers in ML-text and Isar-proofs
Christian Urban <urbanc@in.tum.de>
parents: 113
diff changeset
    58
lemma %linenos trcl_induct:
218
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    59
assumes "trcl R x y"
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    60
shows "(\<And>x. P x x) \<Longrightarrow> (\<And>x y z. R x y \<Longrightarrow> P y z \<Longrightarrow> P x z) \<Longrightarrow> P x y"
113
9b6c9d172378 slightly updated
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    61
apply(atomize (full))
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
    62
apply(cut_tac prems)
113
9b6c9d172378 slightly updated
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    63
apply(unfold trcl_def)
9b6c9d172378 slightly updated
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    64
apply(drule spec[where x=P])
9b6c9d172378 slightly updated
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    65
apply(assumption)
9b6c9d172378 slightly updated
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    66
done
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    67
113
9b6c9d172378 slightly updated
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    68
text {*
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
    69
  The proofs for the introduction rules are slightly more complicated. 
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
    70
  For the first one, we need to prove the following lemma:
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
    71
*}
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
    72
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
    73
lemma %linenos trcl_base: 
218
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    74
shows "trcl R x x"
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
    75
apply(unfold trcl_def)
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
    76
apply(rule allI impI)+
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
    77
apply(drule spec)
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
    78
apply(assumption)
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
    79
done
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
    80
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
    81
text {*
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
    82
  We again unfold first the definition and apply introduction rules 
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
    83
  for @{text "\<forall>"} and @{text "\<longrightarrow>"} as often as possible (Lines 3 and 4).
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 115
diff changeset
    84
  We then end up in the goal state:
113
9b6c9d172378 slightly updated
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    85
*}
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    86
113
9b6c9d172378 slightly updated
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    87
(*<*)lemma "trcl R x x"
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
    88
apply (unfold trcl_def)
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
    89
apply (rule allI impI)+(*>*)
113
9b6c9d172378 slightly updated
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    90
txt {* @{subgoals [display]} *}
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
    91
(*<*)oops(*>*)
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
    92
113
9b6c9d172378 slightly updated
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    93
text {*
218
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    94
  The two assumptions come from the definition of @{term trcl} and correspond
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    95
  to the introduction rules. Thus, all we have to do is to eliminate the
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
    96
  universal quantifier in front of the first assumption (Line 5), and then
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
    97
  solve the goal by @{text assumption} (Line 6).
113
9b6c9d172378 slightly updated
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    98
*}
9b6c9d172378 slightly updated
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    99
9b6c9d172378 slightly updated
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   100
text {*
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   101
  Next we have to show that the second introduction rule also follows from the
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   102
  definition.  Since this rule has premises, the proof is a bit more
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   103
  involved. After unfolding the definitions and applying the introduction
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   104
  rules for @{text "\<forall>"} and @{text "\<longrightarrow>"}
113
9b6c9d172378 slightly updated
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   105
*}
9b6c9d172378 slightly updated
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   106
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   107
lemma trcl_step: 
218
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   108
shows "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   109
apply (unfold trcl_def)
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   110
apply (rule allI impI)+
113
9b6c9d172378 slightly updated
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   111
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   112
txt {* 
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   113
  we obtain the goal state
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   114
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   115
  @{subgoals [display]} 
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   116
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   117
  To see better where we are, let us explicitly name the assumptions 
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   118
  by starting a subproof.
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   119
*}
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   120
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   121
proof -
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   122
  case (goal1 P)
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   123
  have p1: "R x y" by fact
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   124
  have p2: "\<forall>P. (\<forall>x. P x x) 
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   125
                  \<longrightarrow> (\<forall>x y z. R x y \<longrightarrow> P y z \<longrightarrow> P x z) \<longrightarrow> P y z" by fact
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   126
  have r1: "\<forall>x. P x x" by fact
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   127
  have r2: "\<forall>x y z. R x y \<longrightarrow> P y z \<longrightarrow> P x z" by fact
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   128
  show "P x z"
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   129
  
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   130
txt {*
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   131
  The assumptions @{text "p1"} and @{text "p2"} correspond to the premises of
218
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   132
  the second introduction rule (unfolded); the assumptions @{text "r1"} and @{text "r2"}
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   133
  come from the definition of @{term trcl}. We apply @{text "r2"} to the goal
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   134
  @{term "P x z"}. In order for this assumption to be applicable as a rule, we
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   135
  have to eliminate the universal quantifier and turn the object-level
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   136
  implications into meta-level ones. This can be accomplished using the @{text
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   137
  rule_format} attribute. So we continue the proof with:
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   138
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   139
*}
113
9b6c9d172378 slightly updated
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   140
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   141
    apply (rule r2[rule_format])
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   142
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   143
txt {*
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   144
  This gives us two new subgoals
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   145
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   146
  @{subgoals [display]} 
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   147
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   148
  which can be solved using assumptions @{text p1} and @{text p2}. The latter
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   149
  involves a quantifier and implications that have to be eliminated before it
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   150
  can be applied. To avoid potential problems with higher-order unification,
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   151
  we explicitly instantiate the quantifier to @{text "P"} and also match
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   152
  explicitly the implications with @{text "r1"} and @{text "r2"}. This gives
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   153
  the proof:
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   154
*}
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   155
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   156
    apply(rule p1)
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   157
    apply(rule p2[THEN spec[where x=P], THEN mp, THEN mp, OF r1, OF r2])
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   158
    done
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   159
qed
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   160
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   161
text {*
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   162
  Now we are done. It might be surprising that we are not using the automatic
244
dc95a56b1953 fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents: 224
diff changeset
   163
  tactics available in Isabelle/HOL for proving this lemmas. After all @{text
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   164
  "blast"} would easily dispense of it.
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   165
*}
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   166
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   167
lemma trcl_step_blast: 
218
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   168
shows "R x y \<Longrightarrow> trcl R y z \<Longrightarrow> trcl R x z"
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   169
apply(unfold trcl_def)
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   170
apply(blast)
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   171
done
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   172
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   173
text {*
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 115
diff changeset
   174
  Experience has shown that it is generally a bad idea to rely heavily on
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 115
diff changeset
   175
  @{text blast}, @{text auto} and the like in automated proofs. The reason is
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 115
diff changeset
   176
  that you do not have precise control over them (the user can, for example,
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 115
diff changeset
   177
  declare new intro- or simplification rules that can throw automatic tactics
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 115
diff changeset
   178
  off course) and also it is very hard to debug proofs involving automatic
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 115
diff changeset
   179
  tactics whenever something goes wrong. Therefore if possible, automatic 
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
   180
  tactics in packages should be avoided or be constrained sufficiently.
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   181
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 115
diff changeset
   182
  The method of defining inductive predicates by impredicative quantification
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 115
diff changeset
   183
  also generalises to mutually inductive predicates. The next example defines
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   184
  the predicates @{text even} and @{text odd} given by
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   185
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   186
  \begin{center}\small
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   187
  @{prop[mode=Axiom] "even (0::nat)"} \hspace{5mm}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   188
  @{prop[mode=Rule] "odd n \<Longrightarrow> even (Suc n)"} \hspace{5mm}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   189
  @{prop[mode=Rule] "even n \<Longrightarrow> odd (Suc n)"}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   190
  \end{center}
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   191
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   192
  Since the predicates @{term even} and @{term odd} are mutually inductive, each 
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   193
  corresponding definition must quantify over both predicates (we name them 
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   194
  below @{text "P"} and @{text "Q"}).
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   195
*}
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   196
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   197
definition "even \<equiv> 
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   198
  \<lambda>n. \<forall>P Q. P 0 \<longrightarrow> (\<forall>m. Q m \<longrightarrow> P (Suc m)) 
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   199
                 \<longrightarrow> (\<forall>m. P m \<longrightarrow> Q (Suc m)) \<longrightarrow> P n"
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   200
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   201
definition "odd \<equiv>
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   202
  \<lambda>n. \<forall>P Q. P 0 \<longrightarrow> (\<forall>m. Q m \<longrightarrow> P (Suc m)) 
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   203
                 \<longrightarrow> (\<forall>m. P m \<longrightarrow> Q (Suc m)) \<longrightarrow> Q n"
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   204
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   205
text {*
116
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 115
diff changeset
   206
  For proving the induction principles, we use exactly the same technique 
c9ff326e3ce5 more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents: 115
diff changeset
   207
  as in the transitive closure example, namely:
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   208
*}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   209
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   210
lemma even_induct:
218
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   211
assumes "even n"
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   212
shows "P 0 \<Longrightarrow> (\<And>m. Q m \<Longrightarrow> P (Suc m)) \<Longrightarrow> (\<And>m. P m \<Longrightarrow> Q (Suc m)) \<Longrightarrow> P n"
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   213
apply(atomize (full))
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
   214
apply(cut_tac prems)
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   215
apply(unfold even_def)
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   216
apply(drule spec[where x=P])
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   217
apply(drule spec[where x=Q])
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   218
apply(assumption)
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   219
done
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   220
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   221
text {*
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   222
  The only difference with the proof @{text "trcl_induct"} is that we have to
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   223
  instantiate here two universal quantifiers.  We omit the other induction
218
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   224
  principle that has @{prop "even n"} as premise and @{term "Q n"} as conclusion.  
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   225
  The proofs of the introduction rules are also very similar to the ones in 
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   226
  the @{text "trcl"}-example. We only show the proof of the second introduction 
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   227
  rule.
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   228
*}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   229
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   230
lemma %linenos evenS: 
218
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   231
shows "odd m \<Longrightarrow> even (Suc m)"
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   232
apply (unfold odd_def even_def)
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   233
apply (rule allI impI)+
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   234
proof -
165
890fbfef6d6b partially adapted to new antiquotation infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 129
diff changeset
   235
  case (goal1 P Q)
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   236
  have p1: "\<forall>P Q. P 0 \<longrightarrow> (\<forall>m. Q m \<longrightarrow> P (Suc m)) 
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   237
                             \<longrightarrow> (\<forall>m. P m \<longrightarrow> Q (Suc m)) \<longrightarrow> Q m" by fact
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   238
  have r1: "P 0" by fact
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   239
  have r2: "\<forall>m. Q m \<longrightarrow> P (Suc m)" by fact
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   240
  have r3: "\<forall>m. P m \<longrightarrow> Q (Suc m)" by fact
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   241
  show "P (Suc m)"
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   242
    apply(rule r2[rule_format])
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   243
    apply(rule p1[THEN spec[where x=P], THEN spec[where x=Q],
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   244
	           THEN mp, THEN mp, THEN mp, OF r1, OF r2, OF r3])
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   245
    done
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   246
qed
88
ebbd0dd008c8 adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents: 38
diff changeset
   247
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   248
text {*
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
   249
  The interesting lines are 7 to 15. Again, the assumptions fall into two categories:
218
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   250
  @{text p1} corresponds to the premise of the introduction rule; @{text "r1"}
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   251
  to @{text "r3"} come from the definition of @{text "even"}.
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   252
  In Line 13, we apply the assumption @{text "r2"} (since we prove the second
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   253
  introduction rule). In Lines 14 and 15 we apply assumption @{text "p1"} (if
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   254
  the second introduction rule had more premises we have to do that for all
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   255
  of them). In order for this assumption to be applicable, the quantifiers
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   256
  need to be instantiated and then also the implications need to be resolved
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   257
  with the other rules.
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   258
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   259
  Next we define the accessible part of a relation @{text R} given by
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   260
  the single rule:
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   261
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   262
  \begin{center}\small
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   263
  \mbox{\inferrule{@{term "\<And>y. R y x \<Longrightarrow> accpart R y"}}{@{term "accpart R x"}}}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   264
  \end{center}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   265
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
   266
  The interesting point of this definition is that it contains a quantification
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
   267
  that ranges only over the premise and the premise has also a precondition.
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   268
  The definition of @{text "accpart"} is:
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   269
*}
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   270
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   271
definition "accpart \<equiv> \<lambda>R x. \<forall>P. (\<forall>x. (\<forall>y. R y x \<longrightarrow> P y) \<longrightarrow> P x) \<longrightarrow> P x"
32
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   272
5bb2d29553c2 Added new chapter about writing packages.
berghofe
parents:
diff changeset
   273
text {*
218
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   274
  The proof of the induction principle is again straightforward and omitted.
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   275
  Proving the introduction rule is a little more complicated, because the 
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   276
  quantifier and the implication in the premise. The proof is as follows.
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   277
*}
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   278
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   279
lemma %linenos accpartI: 
218
7ff7325e3b4e started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   280
shows "(\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x"
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   281
apply (unfold accpart_def)
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   282
apply (rule allI impI)+
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   283
proof -
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   284
  case (goal1 P)
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   285
  have p1: "\<And>y. R y x \<Longrightarrow> 
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   286
                   (\<forall>P. (\<forall>x. (\<forall>y. R y x \<longrightarrow> P y) \<longrightarrow> P x) \<longrightarrow> P y)" by fact
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   287
  have r1: "\<forall>x. (\<forall>y. R y x \<longrightarrow> P y) \<longrightarrow> P x" by fact
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   288
  show "P x"
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   289
    apply(rule r1[rule_format])
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   290
    proof -
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   291
      case (goal1 y)
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   292
      have r1_prem: "R y x" by fact
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   293
      show "P y"
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   294
	apply(rule p1[OF r1_prem, THEN spec[where x=P], THEN mp, OF r1])
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   295
      done
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   296
  qed
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   297
qed
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   298
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   299
text {*
295
24c68350d059 polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents: 244
diff changeset
   300
  As you can see, there are now two subproofs. The assumptions fall as usual into
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   301
  two categories (Lines 7 to 9). In Line 11, applying the assumption @{text
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   302
  "r1"} generates a goal state with the new local assumption @{term "R y x"},
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   303
  named @{text "r1_prem"} in the second subproof (Line 14). This local assumption is
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   304
  used to solve the goal @{term "P y"} with the help of assumption @{text
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   305
  "p1"}.
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   306
124
0b9fa606a746 added to the first-steps section
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   307
219
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   308
  \begin{exercise}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   309
  Give the definition for the freshness predicate for lambda-terms. The rules
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   310
  for this predicate are:
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   311
  
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   312
  \begin{center}\small
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   313
  @{prop[mode=Rule] "a\<noteq>b \<Longrightarrow> fresh a (Var b)"}\hspace{5mm}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   314
  @{prop[mode=Rule] "\<lbrakk>fresh a t; fresh a s\<rbrakk> \<Longrightarrow> fresh a (App t s)"}\\[2mm]
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   315
  @{prop[mode=Axiom] "fresh a (Lam a t)"}\hspace{5mm}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   316
  @{prop[mode=Rule] "\<lbrakk>a\<noteq>b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)"}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   317
  \end{center}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   318
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   319
  From the definition derive the induction principle and the introduction 
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   320
  rules. 
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   321
  \end{exercise}
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   322
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   323
  The point of all these examples is to get a feeling what the automatic
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   324
  proofs should do in order to solve all inductive definitions we throw at
98d43270024f more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 218
diff changeset
   325
  them.  This is usually the first step in writing a package. We next explain
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   326
  the parsing and typing part of the package.
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   327
115
039845fc96bd some update of the package introduction
Christian Urban <urbanc@in.tum.de>
parents: 114
diff changeset
   328
*}
129
e0d368a45537 started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents: 127
diff changeset
   329
(*<*)end(*>*)