src/Doc/Isar_Ref/Framework.thy
author wenzelm
Sat, 07 Nov 2015 12:53:22 +0100
changeset 61597 53e32a9b66b8
parent 61493 0debd22f0c0e
child 61656 cfabbc083977
permissions -rw-r--r--
added @{undefined} with somewhat undefined symbol;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
29716
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
     1
theory Framework
42651
e3fdb7c96be5 formal Base theory;
wenzelm
parents: 42626
diff changeset
     2
imports Base Main
29716
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
     3
begin
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
     4
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
     5
chapter \<open>The Isabelle/Isar Framework \label{ch:isar-framework}\<close>
29716
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
     6
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
     7
text \<open>
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 56594
diff changeset
     8
  Isabelle/Isar @{cite "Wenzel:1999:TPHOL" and "Wenzel-PhD" and
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 56594
diff changeset
     9
  "Nipkow-TYPES02" and "Wenzel-Paulson:2006" and "Wenzel:2006:Festschrift"}
29716
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    10
  is intended as a generic framework for developing formal
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    11
  mathematical documents with full proof checking.  Definitions and
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
    12
  proofs are organized as theories.  An assembly of theory sources may
29716
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    13
  be presented as a printed document; see also
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    14
  \chref{ch:document-prep}.
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    15
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    16
  The main objective of Isar is the design of a human-readable
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    17
  structured proof language, which is called the ``primary proof
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    18
  format'' in Isar terminology.  Such a primary proof language is
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    19
  somewhere in the middle between the extremes of primitive proof
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    20
  objects and actual natural language.  In this respect, Isar is a bit
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 56594
diff changeset
    21
  more formalistic than Mizar @{cite "Trybulec:1993:MizarFeatures" and
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 56594
diff changeset
    22
  "Rudnicki:1992:MizarOverview" and "Wiedijk:1999:Mizar"},
29716
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    23
  using logical symbols for certain reasoning schemes where Mizar
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 56594
diff changeset
    24
  would prefer English words; see @{cite "Wenzel-Wiedijk:2002"} for
29716
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    25
  further comparisons of these systems.
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    26
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    27
  So Isar challenges the traditional way of recording informal proofs
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    28
  in mathematical prose, as well as the common tendency to see fully
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    29
  formal proofs directly as objects of some logical calculus (e.g.\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    30
  \<open>\<lambda>\<close>-terms in a version of type theory).  In fact, Isar is
29716
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    31
  better understood as an interpreter of a simple block-structured
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
    32
  language for describing the data flow of local facts and goals,
29716
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    33
  interspersed with occasional invocations of proof methods.
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    34
  Everything is reduced to logical inferences internally, but these
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    35
  steps are somewhat marginal compared to the overall bookkeeping of
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    36
  the interpretation process.  Thanks to careful design of the syntax
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    37
  and semantics of Isar language elements, a formal record of Isar
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    38
  instructions may later appear as an intelligible text to the
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    39
  attentive reader.
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    40
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    41
  The Isar proof language has emerged from careful analysis of some
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    42
  inherent virtues of the existing logical framework of Isabelle/Pure
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 56594
diff changeset
    43
  @{cite "paulson-found" and "paulson700"}, notably composition of higher-order
29716
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    44
  natural deduction rules, which is a generalization of Gentzen's
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 56594
diff changeset
    45
  original calculus @{cite "Gentzen:1935"}.  The approach of generic
29716
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    46
  inference systems in Pure is continued by Isar towards actual proof
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    47
  texts.
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    48
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    49
  Concrete applications require another intermediate layer: an
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 56594
diff changeset
    50
  object-logic.  Isabelle/HOL @{cite "isa-tutorial"} (simply-typed
29716
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    51
  set-theory) is being used most of the time; Isabelle/ZF
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 56594
diff changeset
    52
  @{cite "isabelle-ZF"} is less extensively developed, although it would
29716
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
    53
  probably fit better for classical mathematics.
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
    54
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
    55
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
    56
  In order to illustrate natural deduction in Isar, we shall
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
    57
  refer to the background theory and library of Isabelle/HOL.  This
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
    58
  includes common notions of predicate logic, naive set-theory etc.\
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
    59
  using fairly standard mathematical notation.  From the perspective
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
    60
  of generic natural deduction there is nothing special about the
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    61
  logical connectives of HOL (\<open>\<and>\<close>, \<open>\<or>\<close>, \<open>\<forall>\<close>,
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    62
  \<open>\<exists>\<close>, etc.), only the resulting reasoning principles are
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
    63
  relevant to the user.  There are similar rules available for
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    64
  set-theory operators (\<open>\<inter>\<close>, \<open>\<union>\<close>, \<open>\<Inter>\<close>, \<open>\<Union>\<close>, etc.), or any other theory developed in the library (lattice
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
    65
  theory, topology etc.).
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
    66
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
    67
  Subsequently we briefly review fragments of Isar proof texts
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
    68
  corresponding directly to such general deduction schemes.  The
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
    69
  examples shall refer to set-theory, to minimize the danger of
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
    70
  understanding connectives of predicate logic as something special.
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
    71
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
    72
  \<^medskip>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    73
  The following deduction performs \<open>\<inter>\<close>-introduction,
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
    74
  working forwards from assumptions towards the conclusion.  We give
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
    75
  both the Isar text, and depict the primitive rule involved, as
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
    76
  determined by unification of the problem against rules that are
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
    77
  declared in the library context.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
    78
\<close>
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
    79
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
    80
text_raw \<open>\medskip\begin{minipage}{0.6\textwidth}\<close>
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
    81
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
    82
(*<*)
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
    83
notepad
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
    84
begin
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
    85
(*>*)
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
    86
    assume "x \<in> A" and "x \<in> B"
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
    87
    then have "x \<in> A \<inter> B" ..
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
    88
(*<*)
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
    89
end
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
    90
(*>*)
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
    91
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
    92
text_raw \<open>\end{minipage}\begin{minipage}{0.4\textwidth}\<close>
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
    93
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
    94
text \<open>
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
    95
  \infer{@{prop "x \<in> A \<inter> B"}}{@{prop "x \<in> A"} & @{prop "x \<in> B"}}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
    96
\<close>
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
    97
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
    98
text_raw \<open>\end{minipage}\<close>
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
    99
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   100
text \<open>
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   101
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   102
  Note that @{command assume} augments the proof
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   103
  context, @{command then} indicates that the current fact shall be
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   104
  used in the next step, and @{command have} states an intermediate
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   105
  goal.  The two dots ``@{command ".."}'' refer to a complete proof of
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   106
  this claim, using the indicated facts and a canonical rule from the
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   107
  context.  We could have been more explicit here by spelling out the
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   108
  final proof step via the @{command "by"} command:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   109
\<close>
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   110
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   111
(*<*)
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
   112
notepad
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
   113
begin
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   114
(*>*)
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   115
    assume "x \<in> A" and "x \<in> B"
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   116
    then have "x \<in> A \<inter> B" by (rule IntI)
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   117
(*<*)
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
   118
end
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   119
(*>*)
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   120
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   121
text \<open>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   122
  The format of the \<open>\<inter>\<close>-introduction rule represents
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   123
  the most basic inference, which proceeds from given premises to a
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   124
  conclusion, without any nested proof context involved.
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   125
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   126
  The next example performs backwards introduction on @{term "\<Inter>\<A>"},
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   127
  the intersection of all sets within a given set.  This requires a
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   128
  nested proof of set membership within a local context, where @{term
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   129
  A} is an arbitrary-but-fixed member of the collection:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   130
\<close>
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   131
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   132
text_raw \<open>\medskip\begin{minipage}{0.6\textwidth}\<close>
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   133
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   134
(*<*)
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
   135
notepad
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
   136
begin
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   137
(*>*)
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   138
    have "x \<in> \<Inter>\<A>"
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   139
    proof
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   140
      fix A
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   141
      assume "A \<in> \<A>"
29733
f38ccabb2edc improved sorry/noproof markup;
wenzelm
parents: 29732
diff changeset
   142
      show "x \<in> A" sorry %noproof
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   143
    qed
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   144
(*<*)
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
   145
end
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   146
(*>*)
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   147
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   148
text_raw \<open>\end{minipage}\begin{minipage}{0.4\textwidth}\<close>
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   149
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   150
text \<open>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   151
  \infer{@{prop "x \<in> \<Inter>\<A>"}}{\infer*{@{prop "x \<in> A"}}{\<open>[A][A \<in> \<A>]\<close>}}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   152
\<close>
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   153
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   154
text_raw \<open>\end{minipage}\<close>
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   155
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   156
text \<open>
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   157
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   158
  This Isar reasoning pattern again refers to the
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   159
  primitive rule depicted above.  The system determines it in the
56594
e3a06699a13f tuned spelling;
wenzelm
parents: 56451
diff changeset
   160
  ``@{command proof}'' step, which could have been spelled out more
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   161
  explicitly as ``@{command proof}~\<open>(rule InterI)\<close>''.  Note
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   162
  that the rule involves both a local parameter @{term "A"} and an
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   163
  assumption @{prop "A \<in> \<A>"} in the nested reasoning.  This kind of
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   164
  compound rule typically demands a genuine sub-proof in Isar, working
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   165
  backwards rather than forwards as seen before.  In the proof body we
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   166
  encounter the @{command fix}-@{command assume}-@{command show}
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   167
  outline of nested sub-proofs that is typical for Isar.  The final
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   168
  @{command show} is like @{command have} followed by an additional
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   169
  refinement of the enclosing claim, using the rule derived from the
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   170
  proof body.
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   171
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   172
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   173
  The next example involves @{term "\<Union>\<A>"}, which can be
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   174
  characterized as the set of all @{term "x"} such that @{prop "\<exists>A. x
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   175
  \<in> A \<and> A \<in> \<A>"}.  The elimination rule for @{prop "x \<in> \<Union>\<A>"} does
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   176
  not mention \<open>\<exists>\<close> and \<open>\<and>\<close> at all, but admits to obtain
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   177
  directly a local @{term "A"} such that @{prop "x \<in> A"} and @{prop "A
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   178
  \<in> \<A>"} hold.  This corresponds to the following Isar proof and
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   179
  inference rule, respectively:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   180
\<close>
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   181
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   182
text_raw \<open>\medskip\begin{minipage}{0.6\textwidth}\<close>
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   183
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   184
(*<*)
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
   185
notepad
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
   186
begin
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   187
(*>*)
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   188
    assume "x \<in> \<Union>\<A>"
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   189
    then have C
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   190
    proof
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   191
      fix A
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   192
      assume "x \<in> A" and "A \<in> \<A>"
29733
f38ccabb2edc improved sorry/noproof markup;
wenzelm
parents: 29732
diff changeset
   193
      show C sorry %noproof
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   194
    qed
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   195
(*<*)
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
   196
end
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   197
(*>*)
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   198
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   199
text_raw \<open>\end{minipage}\begin{minipage}{0.4\textwidth}\<close>
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   200
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   201
text \<open>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   202
  \infer{@{prop "C"}}{@{prop "x \<in> \<Union>\<A>"} & \infer*{@{prop "C"}~}{\<open>[A][x \<in> A, A \<in> \<A>]\<close>}}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   203
\<close>
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   204
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   205
text_raw \<open>\end{minipage}\<close>
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   206
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   207
text \<open>
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   208
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   209
  Although the Isar proof follows the natural
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   210
  deduction rule closely, the text reads not as natural as
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   211
  anticipated.  There is a double occurrence of an arbitrary
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   212
  conclusion @{prop "C"}, which represents the final result, but is
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   213
  irrelevant for now.  This issue arises for any elimination rule
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   214
  involving local parameters.  Isar provides the derived language
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   215
  element @{command obtain}, which is able to perform the same
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   216
  elimination proof more conveniently:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   217
\<close>
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   218
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   219
(*<*)
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
   220
notepad
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
   221
begin
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   222
(*>*)
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   223
    assume "x \<in> \<Union>\<A>"
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   224
    then obtain A where "x \<in> A" and "A \<in> \<A>" ..
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   225
(*<*)
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
   226
end
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   227
(*>*)
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   228
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   229
text \<open>
61420
ee42cba50933 redundant due to \parindent 0pt;
wenzelm
parents: 60618
diff changeset
   230
  Here we avoid to mention the final conclusion @{prop "C"}
29721
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   231
  and return to plain forward reasoning.  The rule involved in the
df4e53d18ebc added introductory examples;
wenzelm
parents: 29716
diff changeset
   232
  ``@{command ".."}'' proof is the same as before.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   233
\<close>
29716
b6266c4c68fe basic setup for chapter "The Isabelle/Isar Framework";
wenzelm
parents:
diff changeset
   234
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   235
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   236
section \<open>The Pure framework \label{sec:framework-pure}\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   237
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   238
text \<open>
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 56594
diff changeset
   239
  The Pure logic @{cite "paulson-found" and "paulson700"} is an intuitionistic
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 56594
diff changeset
   240
  fragment of higher-order logic @{cite "church40"}.  In type-theoretic
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   241
  parlance, there are three levels of \<open>\<lambda>\<close>-calculus with
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   242
  corresponding arrows \<open>\<Rightarrow>\<close>/\<open>\<And>\<close>/\<open>\<Longrightarrow>\<close>:
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   243
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   244
  \<^medskip>
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   245
  \begin{tabular}{ll}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   246
  \<open>\<alpha> \<Rightarrow> \<beta>\<close> & syntactic function space (terms depending on terms) \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   247
  \<open>\<And>x. B(x)\<close> & universal quantification (proofs depending on terms) \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   248
  \<open>A \<Longrightarrow> B\<close> & implication (proofs depending on proofs) \\
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   249
  \end{tabular}
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   250
  \<^medskip>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   251
61420
ee42cba50933 redundant due to \parindent 0pt;
wenzelm
parents: 60618
diff changeset
   252
  Here only the types of syntactic terms, and the
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   253
  propositions of proof terms have been shown.  The \<open>\<lambda>\<close>-structure of proofs can be recorded as an optional feature of
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 56594
diff changeset
   254
  the Pure inference kernel @{cite "Berghofer-Nipkow:2000:TPHOL"}, but
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61421
diff changeset
   255
  the formal system can never depend on them due to \<^emph>\<open>proof
e467ae7aa808 more control symbols;
wenzelm
parents: 61421
diff changeset
   256
  irrelevance\<close>.
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   257
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   258
  On top of this most primitive layer of proofs, Pure implements a
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   259
  generic calculus for nested natural deduction rules, similar to
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 56594
diff changeset
   260
  @{cite "Schroeder-Heister:1984"}.  Here object-logic inferences are
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   261
  internalized as formulae over \<open>\<And>\<close> and \<open>\<Longrightarrow>\<close>.
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   262
  Combining such rule statements may involve higher-order unification
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 56594
diff changeset
   263
  @{cite "paulson-natural"}.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   264
\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   265
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   266
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   267
subsection \<open>Primitive inferences\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   268
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   269
text \<open>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   270
  Term syntax provides explicit notation for abstraction \<open>\<lambda>x ::
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   271
  \<alpha>. b(x)\<close> and application \<open>b a\<close>, while types are usually
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   272
  implicit thanks to type-inference; terms of type \<open>prop\<close> are
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   273
  called propositions.  Logical statements are composed via \<open>\<And>x
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   274
  :: \<alpha>. B(x)\<close> and \<open>A \<Longrightarrow> B\<close>.  Primitive reasoning operates on
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   275
  judgments of the form \<open>\<Gamma> \<turnstile> \<phi>\<close>, with standard introduction
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   276
  and elimination rules for \<open>\<And>\<close> and \<open>\<Longrightarrow>\<close> that refer to
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   277
  fixed parameters \<open>x\<^sub>1, \<dots>, x\<^sub>m\<close> and hypotheses
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   278
  \<open>A\<^sub>1, \<dots>, A\<^sub>n\<close> from the context \<open>\<Gamma>\<close>;
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   279
  the corresponding proof terms are left implicit.  The subsequent
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   280
  inference rules define \<open>\<Gamma> \<turnstile> \<phi>\<close> inductively, relative to a
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   281
  collection of axioms:
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   282
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   283
  \[
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   284
  \infer{\<open>\<turnstile> A\<close>}{(\<open>A\<close> \text{~axiom})}
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   285
  \qquad
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   286
  \infer{\<open>A \<turnstile> A\<close>}{}
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   287
  \]
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   288
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   289
  \[
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   290
  \infer{\<open>\<Gamma> \<turnstile> \<And>x. B(x)\<close>}{\<open>\<Gamma> \<turnstile> B(x)\<close> & \<open>x \<notin> \<Gamma>\<close>}
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   291
  \qquad
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   292
  \infer{\<open>\<Gamma> \<turnstile> B(a)\<close>}{\<open>\<Gamma> \<turnstile> \<And>x. B(x)\<close>}
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   293
  \]
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   294
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   295
  \[
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   296
  \infer{\<open>\<Gamma> - A \<turnstile> A \<Longrightarrow> B\<close>}{\<open>\<Gamma> \<turnstile> B\<close>}
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   297
  \qquad
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   298
  \infer{\<open>\<Gamma>\<^sub>1 \<union> \<Gamma>\<^sub>2 \<turnstile> B\<close>}{\<open>\<Gamma>\<^sub>1 \<turnstile> A \<Longrightarrow> B\<close> & \<open>\<Gamma>\<^sub>2 \<turnstile> A\<close>}
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   299
  \]
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   300
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   301
  Furthermore, Pure provides a built-in equality \<open>\<equiv> :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   302
  prop\<close> with axioms for reflexivity, substitution, extensionality,
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   303
  and \<open>\<alpha>\<beta>\<eta>\<close>-conversion on \<open>\<lambda>\<close>-terms.
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   304
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   305
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   306
  An object-logic introduces another layer on top of Pure,
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   307
  e.g.\ with types \<open>i\<close> for individuals and \<open>o\<close> for
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   308
  propositions, term constants \<open>Trueprop :: o \<Rightarrow> prop\<close> as
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   309
  (implicit) derivability judgment and connectives like \<open>\<and> :: o
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   310
  \<Rightarrow> o \<Rightarrow> o\<close> or \<open>\<forall> :: (i \<Rightarrow> o) \<Rightarrow> o\<close>, and axioms for object-level
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   311
  rules such as \<open>conjI: A \<Longrightarrow> B \<Longrightarrow> A \<and> B\<close> or \<open>allI: (\<And>x. B
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   312
  x) \<Longrightarrow> \<forall>x. B x\<close>.  Derived object rules are represented as theorems of
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   313
  Pure.  After the initial object-logic setup, further axiomatizations
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   314
  are usually avoided; plain definitions and derived principles are
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   315
  used exclusively.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   316
\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   317
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   318
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   319
subsection \<open>Reasoning with rules \label{sec:framework-resolution}\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   320
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   321
text \<open>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   322
  Primitive inferences mostly serve foundational purposes.  The main
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   323
  reasoning mechanisms of Pure operate on nested natural deduction
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   324
  rules expressed as formulae, using \<open>\<And>\<close> to bind local
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   325
  parameters and \<open>\<Longrightarrow>\<close> to express entailment.  Multiple
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   326
  parameters and premises are represented by repeating these
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   327
  connectives in a right-associative manner.
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   328
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   329
  Since \<open>\<And>\<close> and \<open>\<Longrightarrow>\<close> commute thanks to the theorem
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   330
  @{prop "(A \<Longrightarrow> (\<And>x. B x)) \<equiv> (\<And>x. A \<Longrightarrow> B x)"}, we may assume w.l.o.g.\
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   331
  that rule statements always observe the normal form where
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   332
  quantifiers are pulled in front of implications at each level of
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   333
  nesting.  This means that any Pure proposition may be presented as a
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61421
diff changeset
   334
  \<^emph>\<open>Hereditary Harrop Formula\<close> @{cite "Miller:1991"} which is of the
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   335
  form \<open>\<And>x\<^sub>1 \<dots> x\<^sub>m. H\<^sub>1 \<Longrightarrow> \<dots> H\<^sub>n \<Longrightarrow>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   336
  A\<close> for \<open>m, n \<ge> 0\<close>, and \<open>A\<close> atomic, and \<open>H\<^sub>1, \<dots>, H\<^sub>n\<close> being recursively of the same format.
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   337
  Following the convention that outermost quantifiers are implicit,
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   338
  Horn clauses \<open>A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> A\<close> are a special
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   339
  case of this.
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   340
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   341
  For example, \<open>\<inter>\<close>-introduction rule encountered before is
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   342
  represented as a Pure theorem as follows:
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   343
  \[
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   344
  \<open>IntI:\<close>~@{prop "x \<in> A \<Longrightarrow> x \<in> B \<Longrightarrow> x \<in> A \<inter> B"}
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   345
  \]
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   346
61420
ee42cba50933 redundant due to \parindent 0pt;
wenzelm
parents: 60618
diff changeset
   347
  This is a plain Horn clause, since no further nesting on
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   348
  the left is involved.  The general \<open>\<Inter>\<close>-introduction
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   349
  corresponds to a Hereditary Harrop Formula with one additional level
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   350
  of nesting:
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   351
  \[
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   352
  \<open>InterI:\<close>~@{prop "(\<And>A. A \<in> \<A> \<Longrightarrow> x \<in> A) \<Longrightarrow> x \<in> \<Inter>\<A>"}
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   353
  \]
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   354
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   355
  \<^medskip>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   356
  Goals are also represented as rules: \<open>A\<^sub>1 \<Longrightarrow>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   357
  \<dots> A\<^sub>n \<Longrightarrow> C\<close> states that the sub-goals \<open>A\<^sub>1, \<dots>,
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   358
  A\<^sub>n\<close> entail the result \<open>C\<close>; for \<open>n = 0\<close> the
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   359
  goal is finished.  To allow \<open>C\<close> being a rule statement
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   360
  itself, we introduce the protective marker \<open># :: prop \<Rightarrow>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   361
  prop\<close>, which is defined as identity and hidden from the user.  We
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   362
  initialize and finish goal states as follows:
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   363
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   364
  \[
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   365
  \begin{array}{c@ {\qquad}c}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   366
  \infer[(@{inference_def init})]{\<open>C \<Longrightarrow> #C\<close>}{} &
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   367
  \infer[(@{inference_def finish})]{\<open>C\<close>}{\<open>#C\<close>}
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   368
  \end{array}
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   369
  \]
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   370
61420
ee42cba50933 redundant due to \parindent 0pt;
wenzelm
parents: 60618
diff changeset
   371
  Goal states are refined in intermediate proof steps until
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   372
  a finished form is achieved.  Here the two main reasoning principles
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   373
  are @{inference resolution}, for back-chaining a rule against a
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   374
  sub-goal (replacing it by zero or more sub-goals), and @{inference
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   375
  assumption}, for solving a sub-goal (finding a short-circuit with
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   376
  local assumptions).  Below \<open>\<^vec>x\<close> stands for \<open>x\<^sub>1, \<dots>, x\<^sub>n\<close> (\<open>n \<ge> 0\<close>).
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   377
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   378
  \[
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   379
  \infer[(@{inference_def resolution})]
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   380
  {\<open>(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> \<^vec>A (\<^vec>a \<^vec>x))\<vartheta> \<Longrightarrow> C\<vartheta>\<close>}
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   381
  {\begin{tabular}{rl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   382
    \<open>rule:\<close> &
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   383
    \<open>\<^vec>A \<^vec>a \<Longrightarrow> B \<^vec>a\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   384
    \<open>goal:\<close> &
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   385
    \<open>(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> B' \<^vec>x) \<Longrightarrow> C\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   386
    \<open>goal unifier:\<close> &
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   387
    \<open>(\<lambda>\<^vec>x. B (\<^vec>a \<^vec>x))\<vartheta> = B'\<vartheta>\<close> \\
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   388
   \end{tabular}}
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   389
  \]
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   390
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   391
  \<^medskip>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   392
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   393
  \[
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   394
  \infer[(@{inference_def assumption})]{\<open>C\<vartheta>\<close>}
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   395
  {\begin{tabular}{rl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   396
    \<open>goal:\<close> &
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   397
    \<open>(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> A \<^vec>x) \<Longrightarrow> C\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   398
    \<open>assm unifier:\<close> & \<open>A\<vartheta> = H\<^sub>i\<vartheta>\<close>~~\text{(for some~\<open>H\<^sub>i\<close>)} \\
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   399
   \end{tabular}}
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   400
  \]
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   401
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   402
  The following trace illustrates goal-oriented reasoning in
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   403
  Isabelle/Pure:
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   404
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   405
  {\footnotesize
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   406
  \<^medskip>
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   407
  \begin{tabular}{r@ {\quad}l}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   408
  \<open>(A \<and> B \<Longrightarrow> B \<and> A) \<Longrightarrow> #(A \<and> B \<Longrightarrow> B \<and> A)\<close> & \<open>(init)\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   409
  \<open>(A \<and> B \<Longrightarrow> B) \<Longrightarrow> (A \<and> B \<Longrightarrow> A) \<Longrightarrow> #\<dots>\<close> & \<open>(resolution B \<Longrightarrow> A \<Longrightarrow> B \<and> A)\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   410
  \<open>(A \<and> B \<Longrightarrow> A \<and> B) \<Longrightarrow> (A \<and> B \<Longrightarrow> A) \<Longrightarrow> #\<dots>\<close> & \<open>(resolution A \<and> B \<Longrightarrow> B)\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   411
  \<open>(A \<and> B \<Longrightarrow> A) \<Longrightarrow> #\<dots>\<close> & \<open>(assumption)\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   412
  \<open>(A \<and> B \<Longrightarrow> A \<and> B) \<Longrightarrow> #\<dots>\<close> & \<open>(resolution A \<and> B \<Longrightarrow> A)\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   413
  \<open>#\<dots>\<close> & \<open>(assumption)\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   414
  \<open>A \<and> B \<Longrightarrow> B \<and> A\<close> & \<open>(finish)\<close> \\
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   415
  \end{tabular}
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   416
  \<^medskip>
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   417
  }
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   418
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   419
  Compositions of @{inference assumption} after @{inference
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   420
  resolution} occurs quite often, typically in elimination steps.
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   421
  Traditional Isabelle tactics accommodate this by a combined
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   422
  @{inference_def elim_resolution} principle.  In contrast, Isar uses
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   423
  a slightly more refined combination, where the assumptions to be
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   424
  closed are marked explicitly, using again the protective marker
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   425
  \<open>#\<close>:
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   426
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   427
  \[
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   428
  \infer[(@{inference refinement})]
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   429
  {\<open>(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> \<^vec>G' (\<^vec>a \<^vec>x))\<vartheta> \<Longrightarrow> C\<vartheta>\<close>}
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   430
  {\begin{tabular}{rl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   431
    \<open>sub\<hyphen>proof:\<close> &
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   432
    \<open>\<^vec>G \<^vec>a \<Longrightarrow> B \<^vec>a\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   433
    \<open>goal:\<close> &
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   434
    \<open>(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> B' \<^vec>x) \<Longrightarrow> C\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   435
    \<open>goal unifier:\<close> &
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   436
    \<open>(\<lambda>\<^vec>x. B (\<^vec>a \<^vec>x))\<vartheta> = B'\<vartheta>\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   437
    \<open>assm unifiers:\<close> &
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   438
    \<open>(\<lambda>\<^vec>x. G\<^sub>j (\<^vec>a \<^vec>x))\<vartheta> = #H\<^sub>i\<vartheta>\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   439
    & \quad (for each marked \<open>G\<^sub>j\<close> some \<open>#H\<^sub>i\<close>) \\
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   440
   \end{tabular}}
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   441
  \]
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   442
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   443
  Here the \<open>sub\<hyphen>proof\<close> rule stems from the
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   444
  main @{command fix}-@{command assume}-@{command show} outline of
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   445
  Isar (cf.\ \secref{sec:framework-subproof}): each assumption
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   446
  indicated in the text results in a marked premise \<open>G\<close> above.
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   447
  The marking enforces resolution against one of the sub-goal's
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   448
  premises.  Consequently, @{command fix}-@{command assume}-@{command
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   449
  show} enables to fit the result of a sub-proof quite robustly into a
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   450
  pending sub-goal, while maintaining a good measure of flexibility.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   451
\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   452
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   453
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   454
section \<open>The Isar proof language \label{sec:framework-isar}\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   455
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   456
text \<open>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   457
  Structured proofs are presented as high-level expressions for
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   458
  composing entities of Pure (propositions, facts, and goals).  The
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   459
  Isar proof language allows to organize reasoning within the
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   460
  underlying rule calculus of Pure, but Isar is not another logical
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   461
  calculus!
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   462
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   463
  Isar is an exercise in sound minimalism.  Approximately half of the
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   464
  language is introduced as primitive, the rest defined as derived
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   465
  concepts.  The following grammar describes the core language
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   466
  (category \<open>proof\<close>), which is embedded into theory
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   467
  specification elements such as @{command theorem}; see also
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   468
  \secref{sec:framework-stmt} for the separate category \<open>statement\<close>.
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   469
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   470
  \<^medskip>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   471
  \begin{tabular}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   472
    \<open>theory\<hyphen>stmt\<close> & = & @{command "theorem"}~\<open>statement proof  |\<close>~~@{command "definition"}~\<open>\<dots>  |  \<dots>\<close> \\[1ex]
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   473
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   474
    \<open>proof\<close> & = & \<open>prfx\<^sup>*\<close>~@{command "proof"}~\<open>method\<^sup>? stmt\<^sup>*\<close>~@{command "qed"}~\<open>method\<^sup>?\<close> \\[1ex]
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   475
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   476
    \<open>prfx\<close> & = & @{command "using"}~\<open>facts\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   477
    & \<open>|\<close> & @{command "unfolding"}~\<open>facts\<close> \\
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   478
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   479
    \<open>stmt\<close> & = & @{command "{"}~\<open>stmt\<^sup>*\<close>~@{command "}"} \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   480
    & \<open>|\<close> & @{command "next"} \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   481
    & \<open>|\<close> & @{command "note"}~\<open>name = facts\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   482
    & \<open>|\<close> & @{command "let"}~\<open>term = term\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   483
    & \<open>|\<close> & @{command "fix"}~\<open>var\<^sup>+\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   484
    & \<open>|\<close> & @{command assume}~\<open>\<guillemotleft>inference\<guillemotright> name: props\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   485
    & \<open>|\<close> & @{command "then"}\<open>\<^sup>?\<close>~\<open>goal\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   486
    \<open>goal\<close> & = & @{command "have"}~\<open>name: props proof\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   487
    & \<open>|\<close> & @{command "show"}~\<open>name: props proof\<close> \\
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   488
  \end{tabular}
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   489
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   490
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   491
  Simultaneous propositions or facts may be separated by the
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   492
  @{keyword "and"} keyword.
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   493
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   494
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   495
  The syntax for terms and propositions is inherited from
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   496
  Pure (and the object-logic).  A \<open>pattern\<close> is a \<open>term\<close> with schematic variables, to be bound by higher-order
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   497
  matching.
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   498
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   499
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   500
  Facts may be referenced by name or proposition.  For
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   501
  example, the result of ``@{command have}~\<open>a: A \<langle>proof\<rangle>\<close>''
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   502
  becomes available both as \<open>a\<close> and
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   503
  \isacharbackquoteopen\<open>A\<close>\isacharbackquoteclose.  Moreover,
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   504
  fact expressions may involve attributes that modify either the
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   505
  theorem or the background context.  For example, the expression
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   506
  ``\<open>a [OF b]\<close>'' refers to the composition of two facts
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   507
  according to the @{inference resolution} inference of
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   508
  \secref{sec:framework-resolution}, while ``\<open>a [intro]\<close>''
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   509
  declares a fact as introduction rule in the context.
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   510
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   511
  The special fact called ``@{fact this}'' always refers to the last
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   512
  result, as produced by @{command note}, @{command assume}, @{command
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   513
  have}, or @{command show}.  Since @{command note} occurs
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   514
  frequently together with @{command then} we provide some
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   515
  abbreviations:
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   516
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   517
  \<^medskip>
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   518
  \begin{tabular}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   519
    @{command from}~\<open>a\<close> & \<open>\<equiv>\<close> & @{command note}~\<open>a\<close>~@{command then} \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   520
    @{command with}~\<open>a\<close> & \<open>\<equiv>\<close> & @{command from}~\<open>a \<AND> this\<close> \\
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   521
  \end{tabular}
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   522
  \<^medskip>
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   523
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   524
  The \<open>method\<close> category is essentially a parameter and may be
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   525
  populated later.  Methods use the facts indicated by @{command
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   526
  "then"} or @{command using}, and then operate on the goal state.
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   527
  Some basic methods are predefined: ``@{method "-"}'' leaves the goal
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   528
  unchanged, ``@{method this}'' applies the facts as rules to the
42626
6ac8c55c657e eliminated some duplicate "def" positions;
wenzelm
parents: 40964
diff changeset
   529
  goal, ``@{method (Pure) "rule"}'' applies the facts to another rule and the
6ac8c55c657e eliminated some duplicate "def" positions;
wenzelm
parents: 40964
diff changeset
   530
  result to the goal (both ``@{method this}'' and ``@{method (Pure) rule}''
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   531
  refer to @{inference resolution} of
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   532
  \secref{sec:framework-resolution}).  The secondary arguments to
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   533
  ``@{method (Pure) rule}'' may be specified explicitly as in ``\<open>(rule
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   534
  a)\<close>'', or picked from the context.  In the latter case, the system
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   535
  first tries rules declared as @{attribute (Pure) elim} or
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   536
  @{attribute (Pure) dest}, followed by those declared as @{attribute
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   537
  (Pure) intro}.
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   538
60618
4c79543cc376 renamed "default" to "standard", to make semantically clear what it is;
wenzelm
parents: 58999
diff changeset
   539
  The default method for @{command proof} is ``@{method standard}''
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   540
  (arguments picked from the context), for @{command qed} it is
60618
4c79543cc376 renamed "default" to "standard", to make semantically clear what it is;
wenzelm
parents: 58999
diff changeset
   541
  ``@{method "succeed"}''.  Further abbreviations for terminal proof steps
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   542
  are ``@{command "by"}~\<open>method\<^sub>1 method\<^sub>2\<close>'' for
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   543
  ``@{command proof}~\<open>method\<^sub>1\<close>~@{command qed}~\<open>method\<^sub>2\<close>'', and ``@{command ".."}'' for ``@{command
60618
4c79543cc376 renamed "default" to "standard", to make semantically clear what it is;
wenzelm
parents: 58999
diff changeset
   544
  "by"}~@{method standard}, and ``@{command "."}'' for ``@{command
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   545
  "by"}~@{method this}''.  The @{command unfolding} element operates
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   546
  directly on the current facts and goal by applying equalities.
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   547
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   548
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   549
  Block structure can be indicated explicitly by ``@{command
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   550
  "{"}~\<open>\<dots>\<close>~@{command "}"}'', although the body of a sub-proof
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   551
  already involves implicit nesting.  In any case, @{command next}
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   552
  jumps into the next section of a block, i.e.\ it acts like closing
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   553
  an implicit block scope and opening another one; there is no direct
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   554
  correspondence to subgoals here.
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   555
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   556
  The remaining elements @{command fix} and @{command assume} build up
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   557
  a local context (see \secref{sec:framework-context}), while
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   558
  @{command show} refines a pending sub-goal by the rule resulting
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   559
  from a nested sub-proof (see \secref{sec:framework-subproof}).
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   560
  Further derived concepts will support calculational reasoning (see
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   561
  \secref{sec:framework-calc}).
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   562
\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   563
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   564
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   565
subsection \<open>Context elements \label{sec:framework-context}\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   566
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   567
text \<open>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   568
  In judgments \<open>\<Gamma> \<turnstile> \<phi>\<close> of the primitive framework, \<open>\<Gamma>\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   569
  essentially acts like a proof context.  Isar elaborates this idea
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   570
  towards a higher-level notion, with additional information for
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   571
  type-inference, term abbreviations, local facts, hypotheses etc.
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   572
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   573
  The element @{command fix}~\<open>x :: \<alpha>\<close> declares a local
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   574
  parameter, i.e.\ an arbitrary-but-fixed entity of a given type; in
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   575
  results exported from the context, \<open>x\<close> may become anything.
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   576
  The @{command assume}~\<open>\<guillemotleft>inference\<guillemotright>\<close> element provides a
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   577
  general interface to hypotheses: ``@{command assume}~\<open>\<guillemotleft>inference\<guillemotright> A\<close>'' produces \<open>A \<turnstile> A\<close> locally, while the
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   578
  included inference tells how to discharge \<open>A\<close> from results
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   579
  \<open>A \<turnstile> B\<close> later on.  There is no user-syntax for \<open>\<guillemotleft>inference\<guillemotright>\<close>, i.e.\ it may only occur internally when derived
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   580
  commands are defined in ML.
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   581
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   582
  At the user-level, the default inference for @{command assume} is
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   583
  @{inference discharge} as given below.  The additional variants
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   584
  @{command presume} and @{command def} are defined as follows:
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   585
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   586
  \<^medskip>
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   587
  \begin{tabular}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   588
    @{command presume}~\<open>A\<close> & \<open>\<equiv>\<close> & @{command assume}~\<open>\<guillemotleft>weak\<hyphen>discharge\<guillemotright> A\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   589
    @{command def}~\<open>x \<equiv> a\<close> & \<open>\<equiv>\<close> & @{command fix}~\<open>x\<close>~@{command assume}~\<open>\<guillemotleft>expansion\<guillemotright> x \<equiv> a\<close> \\
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   590
  \end{tabular}
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   591
  \<^medskip>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   592
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   593
  \[
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   594
  \infer[(@{inference_def discharge})]{\<open>\<strut>\<Gamma> - A \<turnstile> #A \<Longrightarrow> B\<close>}{\<open>\<strut>\<Gamma> \<turnstile> B\<close>}
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   595
  \]
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   596
  \[
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   597
  \infer[(@{inference_def "weak\<hyphen>discharge"})]{\<open>\<strut>\<Gamma> - A \<turnstile> A \<Longrightarrow> B\<close>}{\<open>\<strut>\<Gamma> \<turnstile> B\<close>}
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   598
  \]
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   599
  \[
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   600
  \infer[(@{inference_def expansion})]{\<open>\<strut>\<Gamma> - (x \<equiv> a) \<turnstile> B a\<close>}{\<open>\<strut>\<Gamma> \<turnstile> B x\<close>}
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   601
  \]
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   602
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   603
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   604
  Note that @{inference discharge} and @{inference
42666
fee67c099d03 use existing \<hyphen>;
wenzelm
parents: 42651
diff changeset
   605
  "weak\<hyphen>discharge"} differ in the marker for @{prop A}, which is
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   606
  relevant when the result of a @{command fix}-@{command
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   607
  assume}-@{command show} outline is composed with a pending goal,
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   608
  cf.\ \secref{sec:framework-subproof}.
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   609
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   610
  The most interesting derived context element in Isar is @{command
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 56594
diff changeset
   611
  obtain} @{cite \<open>\S5.3\<close> "Wenzel-PhD"}, which supports generalized
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   612
  elimination steps in a purely forward manner.  The @{command obtain}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   613
  command takes a specification of parameters \<open>\<^vec>x\<close> and
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   614
  assumptions \<open>\<^vec>A\<close> to be added to the context, together
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   615
  with a proof of a case rule stating that this extension is
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   616
  conservative (i.e.\ may be removed from closed results later on):
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   617
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   618
  \<^medskip>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   619
  \begin{tabular}{l}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   620
  \<open>\<langle>facts\<rangle>\<close>~~@{command obtain}~\<open>\<^vec>x \<WHERE> \<^vec>A \<^vec>x  \<langle>proof\<rangle> \<equiv>\<close> \\[0.5ex]
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   621
  \quad @{command have}~\<open>case: \<And>thesis. (\<And>\<^vec>x. \<^vec>A \<^vec>x \<Longrightarrow> thesis) \<Longrightarrow> thesis\<rangle>\<close> \\
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   622
  \quad @{command proof}~@{method "-"} \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   623
  \qquad @{command fix}~\<open>thesis\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   624
  \qquad @{command assume}~\<open>[intro]: \<And>\<^vec>x. \<^vec>A \<^vec>x \<Longrightarrow> thesis\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   625
  \qquad @{command show}~\<open>thesis\<close>~@{command using}~\<open>\<langle>facts\<rangle> \<langle>proof\<rangle>\<close> \\
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   626
  \quad @{command qed} \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   627
  \quad @{command fix}~\<open>\<^vec>x\<close>~@{command assume}~\<open>\<guillemotleft>elimination case\<guillemotright> \<^vec>A \<^vec>x\<close> \\
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   628
  \end{tabular}
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   629
  \<^medskip>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   630
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   631
  \[
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   632
  \infer[(@{inference elimination})]{\<open>\<Gamma> \<turnstile> B\<close>}{
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   633
    \begin{tabular}{rl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   634
    \<open>case:\<close> &
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   635
    \<open>\<Gamma> \<turnstile> \<And>thesis. (\<And>\<^vec>x. \<^vec>A \<^vec>x \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close> \\[0.2ex]
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   636
    \<open>result:\<close> &
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   637
    \<open>\<Gamma> \<union> \<^vec>A \<^vec>y \<turnstile> B\<close> \\[0.2ex]
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   638
    \end{tabular}}
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   639
  \]
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   640
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   641
  Here the name ``\<open>thesis\<close>'' is a specific convention
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   642
  for an arbitrary-but-fixed proposition; in the primitive natural
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   643
  deduction rules shown before we have occasionally used \<open>C\<close>.
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   644
  The whole statement of ``@{command obtain}~\<open>x\<close>~@{keyword
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   645
  "where"}~\<open>A x\<close>'' may be read as a claim that \<open>A x\<close>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   646
  may be assumed for some arbitrary-but-fixed \<open>x\<close>.  Also note
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   647
  that ``@{command obtain}~\<open>A \<AND> B\<close>'' without parameters
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   648
  is similar to ``@{command have}~\<open>A \<AND> B\<close>'', but the
29737
866841668584 replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
wenzelm
parents: 29735
diff changeset
   649
  latter involves multiple sub-goals.
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   650
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   651
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   652
  The subsequent Isar proof texts explain all context
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   653
  elements introduced above using the formal proof language itself.
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   654
  After finishing a local proof within a block, we indicate the
29739
wenzelm
parents: 29738
diff changeset
   655
  exported result via @{command note}.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   656
\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   657
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   658
(*<*)
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   659
theorem True
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   660
proof
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   661
(*>*)
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58618
diff changeset
   662
  text_raw \<open>\begin{minipage}[t]{0.45\textwidth}\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   663
  {
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   664
    fix x
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   665
    have "B x" sorry %noproof
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   666
  }
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   667
  note \<open>\<And>x. B x\<close>
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58618
diff changeset
   668
  text_raw \<open>\end{minipage}\quad\begin{minipage}[t]{0.45\textwidth}\<close>(*<*)next(*>*)
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   669
  {
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   670
    assume A
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   671
    have B sorry %noproof
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   672
  }
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   673
  note \<open>A \<Longrightarrow> B\<close>
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58618
diff changeset
   674
  text_raw \<open>\end{minipage}\\[3ex]\begin{minipage}[t]{0.45\textwidth}\<close>(*<*)next(*>*)
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   675
  {
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   676
    def x \<equiv> a
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   677
    have "B x" sorry %noproof
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   678
  }
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   679
  note \<open>B a\<close>
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58618
diff changeset
   680
  text_raw \<open>\end{minipage}\quad\begin{minipage}[t]{0.45\textwidth}\<close>(*<*)next(*>*)
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   681
  {
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   682
    obtain x where "A x" sorry %noproof
29733
f38ccabb2edc improved sorry/noproof markup;
wenzelm
parents: 29732
diff changeset
   683
    have B sorry %noproof
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   684
  }
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   685
  note \<open>B\<close>
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58618
diff changeset
   686
  text_raw \<open>\end{minipage}\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   687
(*<*)
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   688
qed
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   689
(*>*)
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   690
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   691
text \<open>
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   692
  \<^bigskip>
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   693
  This illustrates the meaning of Isar context
29739
wenzelm
parents: 29738
diff changeset
   694
  elements without goals getting in between.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   695
\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   696
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   697
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   698
subsection \<open>Structured statements \label{sec:framework-stmt}\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   699
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   700
text \<open>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   701
  The category \<open>statement\<close> of top-level theorem specifications
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   702
  is defined as follows:
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   703
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   704
  \<^medskip>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   705
  \begin{tabular}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   706
  \<open>statement\<close> & \<open>\<equiv>\<close> & \<open>name: props \<AND> \<dots>\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   707
  & \<open>|\<close> & \<open>context\<^sup>* conclusion\<close> \\[0.5ex]
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   708
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   709
  \<open>context\<close> & \<open>\<equiv>\<close> & \<open>\<FIXES> vars \<AND> \<dots>\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   710
  & \<open>|\<close> & \<open>\<ASSUMES> name: props \<AND> \<dots>\<close> \\
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   711
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   712
  \<open>conclusion\<close> & \<open>\<equiv>\<close> & \<open>\<SHOWS> name: props \<AND> \<dots>\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   713
  & \<open>|\<close> & \<open>\<OBTAINS> vars \<AND> \<dots> \<WHERE> name: props \<AND> \<dots>\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   714
  & & \quad \<open>\<BBAR> \<dots>\<close> \\
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   715
  \end{tabular}
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   716
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   717
  \<^medskip>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   718
  A simple \<open>statement\<close> consists of named
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   719
  propositions.  The full form admits local context elements followed
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   720
  by the actual conclusions, such as ``@{keyword "fixes"}~\<open>x\<close>~@{keyword "assumes"}~\<open>A x\<close>~@{keyword "shows"}~\<open>B
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   721
  x\<close>''.  The final result emerges as a Pure rule after discharging
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   722
  the context: @{prop "\<And>x. A x \<Longrightarrow> B x"}.
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   723
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   724
  The @{keyword "obtains"} variant is another abbreviation defined
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   725
  below; unlike @{command obtain} (cf.\
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   726
  \secref{sec:framework-context}) there may be several ``cases''
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   727
  separated by ``\<open>\<BBAR>\<close>'', each consisting of several
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   728
  parameters (\<open>vars\<close>) and several premises (\<open>props\<close>).
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   729
  This specifies multi-branch elimination rules.
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   730
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   731
  \<^medskip>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   732
  \begin{tabular}{l}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   733
  \<open>\<OBTAINS> \<^vec>x \<WHERE> \<^vec>A \<^vec>x   \<BBAR>   \<dots>   \<equiv>\<close> \\[0.5ex]
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   734
  \quad \<open>\<FIXES> thesis\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   735
  \quad \<open>\<ASSUMES> [intro]: \<And>\<^vec>x. \<^vec>A \<^vec>x \<Longrightarrow> thesis  \<AND>  \<dots>\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   736
  \quad \<open>\<SHOWS> thesis\<close> \\
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   737
  \end{tabular}
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   738
  \<^medskip>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   739
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   740
  Presenting structured statements in such an ``open'' format usually
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   741
  simplifies the subsequent proof, because the outer structure of the
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   742
  problem is already laid out directly.  E.g.\ consider the following
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   743
  canonical patterns for \<open>\<SHOWS>\<close> and \<open>\<OBTAINS>\<close>,
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   744
  respectively:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   745
\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   746
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   747
text_raw \<open>\begin{minipage}{0.5\textwidth}\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   748
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   749
theorem
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   750
  fixes x and y
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   751
  assumes "A x" and "B y"
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   752
  shows "C x y"
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   753
proof -
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   754
  from \<open>A x\<close> and \<open>B y\<close>
29733
f38ccabb2edc improved sorry/noproof markup;
wenzelm
parents: 29732
diff changeset
   755
  show "C x y" sorry %noproof
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   756
qed
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   757
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   758
text_raw \<open>\end{minipage}\begin{minipage}{0.5\textwidth}\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   759
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   760
theorem
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   761
  obtains x and y
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   762
  where "A x" and "B y"
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   763
proof -
29733
f38ccabb2edc improved sorry/noproof markup;
wenzelm
parents: 29732
diff changeset
   764
  have "A a" and "B b" sorry %noproof
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   765
  then show thesis ..
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   766
qed
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   767
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   768
text_raw \<open>\end{minipage}\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   769
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   770
text \<open>
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   771
  \<^medskip>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   772
  Here local facts \isacharbackquoteopen\<open>A
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   773
  x\<close>\isacharbackquoteclose\ and \isacharbackquoteopen\<open>B
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   774
  y\<close>\isacharbackquoteclose\ are referenced immediately; there is no
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   775
  need to decompose the logical rule structure again.  In the second
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   776
  proof the final ``@{command then}~@{command show}~\<open>thesis\<close>~@{command ".."}''  involves the local rule case \<open>\<And>x
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   777
  y. A x \<Longrightarrow> B y \<Longrightarrow> thesis\<close> for the particular instance of terms \<open>a\<close> and \<open>b\<close> produced in the body.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   778
\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   779
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   780
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   781
subsection \<open>Structured proof refinement \label{sec:framework-subproof}\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   782
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   783
text \<open>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   784
  By breaking up the grammar for the Isar proof language, we may
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   785
  understand a proof text as a linear sequence of individual proof
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   786
  commands.  These are interpreted as transitions of the Isar virtual
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   787
  machine (Isar/VM), which operates on a block-structured
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   788
  configuration in single steps.  This allows users to write proof
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   789
  texts in an incremental manner, and inspect intermediate
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   790
  configurations for debugging.
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   791
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   792
  The basic idea is analogous to evaluating algebraic expressions on a
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   793
  stack machine: \<open>(a + b) \<cdot> c\<close> then corresponds to a sequence
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   794
  of single transitions for each symbol \<open>(, a, +, b, ), \<cdot>, c\<close>.
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   795
  In Isar the algebraic values are facts or goals, and the operations
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   796
  are inferences.
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   797
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   798
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   799
  The Isar/VM state maintains a stack of nodes, each node
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   800
  contains the local proof context, the linguistic mode, and a pending
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   801
  goal (optional).  The mode determines the type of transition that
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   802
  may be performed next, it essentially alternates between forward and
29738
05d5615e12d3 added Isar/VM mode transition diagram;
wenzelm
parents: 29737
diff changeset
   803
  backward reasoning, with an intermediate stage for chained facts
05d5615e12d3 added Isar/VM mode transition diagram;
wenzelm
parents: 29737
diff changeset
   804
  (see \figref{fig:isar-vm}).
05d5615e12d3 added Isar/VM mode transition diagram;
wenzelm
parents: 29737
diff changeset
   805
05d5615e12d3 added Isar/VM mode transition diagram;
wenzelm
parents: 29737
diff changeset
   806
  \begin{figure}[htb]
05d5615e12d3 added Isar/VM mode transition diagram;
wenzelm
parents: 29737
diff changeset
   807
  \begin{center}
48958
12afbf6eb7f9 more standard document preparation within session context;
wenzelm
parents: 48279
diff changeset
   808
  \includegraphics[width=0.8\textwidth]{isar-vm}
29738
05d5615e12d3 added Isar/VM mode transition diagram;
wenzelm
parents: 29737
diff changeset
   809
  \end{center}
05d5615e12d3 added Isar/VM mode transition diagram;
wenzelm
parents: 29737
diff changeset
   810
  \caption{Isar/VM modes}\label{fig:isar-vm}
05d5615e12d3 added Isar/VM mode transition diagram;
wenzelm
parents: 29737
diff changeset
   811
  \end{figure}
05d5615e12d3 added Isar/VM mode transition diagram;
wenzelm
parents: 29737
diff changeset
   812
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   813
  For example, in \<open>state\<close> mode Isar acts like a mathematical
29738
05d5615e12d3 added Isar/VM mode transition diagram;
wenzelm
parents: 29737
diff changeset
   814
  scratch-pad, accepting declarations like @{command fix}, @{command
05d5615e12d3 added Isar/VM mode transition diagram;
wenzelm
parents: 29737
diff changeset
   815
  assume}, and claims like @{command have}, @{command show}.  A goal
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   816
  statement changes the mode to \<open>prove\<close>, which means that we
29738
05d5615e12d3 added Isar/VM mode transition diagram;
wenzelm
parents: 29737
diff changeset
   817
  may now refine the problem via @{command unfolding} or @{command
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   818
  proof}.  Then we are again in \<open>state\<close> mode of a proof body,
29738
05d5615e12d3 added Isar/VM mode transition diagram;
wenzelm
parents: 29737
diff changeset
   819
  which may issue @{command show} statements to solve pending
05d5615e12d3 added Isar/VM mode transition diagram;
wenzelm
parents: 29737
diff changeset
   820
  sub-goals.  A concluding @{command qed} will return to the original
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   821
  \<open>state\<close> mode one level upwards.  The subsequent Isar/VM
29738
05d5615e12d3 added Isar/VM mode transition diagram;
wenzelm
parents: 29737
diff changeset
   822
  trace indicates block structure, linguistic mode, goal state, and
05d5615e12d3 added Isar/VM mode transition diagram;
wenzelm
parents: 29737
diff changeset
   823
  inferences:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   824
\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   825
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   826
text_raw \<open>\begingroup\footnotesize\<close>
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
   827
(*<*)notepad begin
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   828
(*>*)
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58618
diff changeset
   829
  text_raw \<open>\begin{minipage}[t]{0.18\textwidth}\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   830
  have "A \<longrightarrow> B"
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   831
  proof
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   832
    assume A
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   833
    show B
29733
f38ccabb2edc improved sorry/noproof markup;
wenzelm
parents: 29732
diff changeset
   834
      sorry %noproof
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   835
  qed
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58618
diff changeset
   836
  text_raw \<open>\end{minipage}\quad
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   837
\begin{minipage}[t]{0.06\textwidth}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   838
\<open>begin\<close> \\
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   839
\\
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   840
\\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   841
\<open>begin\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   842
\<open>end\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   843
\<open>end\<close> \\
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   844
\end{minipage}
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   845
\begin{minipage}[t]{0.08\textwidth}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   846
\<open>prove\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   847
\<open>state\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   848
\<open>state\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   849
\<open>prove\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   850
\<open>state\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   851
\<open>state\<close> \\
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   852
\end{minipage}\begin{minipage}[t]{0.35\textwidth}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   853
\<open>(A \<longrightarrow> B) \<Longrightarrow> #(A \<longrightarrow> B)\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   854
\<open>(A \<Longrightarrow> B) \<Longrightarrow> #(A \<longrightarrow> B)\<close> \\
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   855
\\
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   856
\\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   857
\<open>#(A \<longrightarrow> B)\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   858
\<open>A \<longrightarrow> B\<close> \\
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   859
\end{minipage}\begin{minipage}[t]{0.4\textwidth}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   860
\<open>(init)\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   861
\<open>(resolution impI)\<close> \\
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   862
\\
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   863
\\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   864
\<open>(refinement #A \<Longrightarrow> B)\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   865
\<open>(finish)\<close> \\
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   866
\end{minipage}\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   867
(*<*)
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
   868
end
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   869
(*>*)
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   870
text_raw \<open>\endgroup\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   871
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   872
text \<open>
61420
ee42cba50933 redundant due to \parindent 0pt;
wenzelm
parents: 60618
diff changeset
   873
  Here the @{inference refinement} inference from
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   874
  \secref{sec:framework-resolution} mediates composition of Isar
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   875
  sub-proofs nicely.  Observe that this principle incorporates some
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   876
  degree of freedom in proof composition.  In particular, the proof
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   877
  body allows parameters and assumptions to be re-ordered, or commuted
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   878
  according to Hereditary Harrop Form.  Moreover, context elements
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   879
  that are not used in a sub-proof may be omitted altogether.  For
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   880
  example:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   881
\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   882
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   883
text_raw \<open>\begin{minipage}{0.5\textwidth}\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   884
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   885
(*<*)
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
   886
notepad
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
   887
begin
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   888
(*>*)
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   889
  have "\<And>x y. A x \<Longrightarrow> B y \<Longrightarrow> C x y"
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   890
  proof -
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   891
    fix x and y
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   892
    assume "A x" and "B y"
29733
f38ccabb2edc improved sorry/noproof markup;
wenzelm
parents: 29732
diff changeset
   893
    show "C x y" sorry %noproof
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   894
  qed
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   895
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58618
diff changeset
   896
text_raw \<open>\end{minipage}\begin{minipage}{0.5\textwidth}\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   897
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   898
(*<*)
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   899
next
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   900
(*>*)
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   901
  have "\<And>x y. A x \<Longrightarrow> B y \<Longrightarrow> C x y"
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   902
  proof -
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   903
    fix x assume "A x"
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   904
    fix y assume "B y"
29733
f38ccabb2edc improved sorry/noproof markup;
wenzelm
parents: 29732
diff changeset
   905
    show "C x y" sorry %noproof
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   906
  qed
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   907
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58618
diff changeset
   908
text_raw \<open>\end{minipage}\\[3ex]\begin{minipage}{0.5\textwidth}\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   909
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   910
(*<*)
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   911
next
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   912
(*>*)
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   913
  have "\<And>x y. A x \<Longrightarrow> B y \<Longrightarrow> C x y"
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   914
  proof -
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   915
    fix y assume "B y"
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   916
    fix x assume "A x"
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   917
    show "C x y" sorry
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   918
  qed
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   919
58999
ed09ae4ea2d8 uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
wenzelm
parents: 58618
diff changeset
   920
text_raw \<open>\end{minipage}\begin{minipage}{0.5\textwidth}\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   921
(*<*)
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   922
next
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   923
(*>*)
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   924
  have "\<And>x y. A x \<Longrightarrow> B y \<Longrightarrow> C x y"
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   925
  proof -
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   926
    fix y assume "B y"
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   927
    fix x
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   928
    show "C x y" sorry
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   929
  qed
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   930
(*<*)
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
   931
end
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   932
(*>*)
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   933
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   934
text_raw \<open>\end{minipage}\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   935
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   936
text \<open>
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   937
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   938
  Such ``peephole optimizations'' of Isar texts are
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   939
  practically important to improve readability, by rearranging
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   940
  contexts elements according to the natural flow of reasoning in the
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   941
  body, while still observing the overall scoping rules.
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   942
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   943
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
   944
  This illustrates the basic idea of structured proof
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   945
  processing in Isar.  The main mechanisms are based on natural
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   946
  deduction rule composition within the Pure framework.  In
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   947
  particular, there are no direct operations on goal states within the
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   948
  proof body.  Moreover, there is no hidden automated reasoning
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   949
  involved, just plain unification.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   950
\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   951
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   952
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   953
subsection \<open>Calculational reasoning \label{sec:framework-calc}\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   954
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   955
text \<open>
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   956
  The existing Isar infrastructure is sufficiently flexible to support
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   957
  calculational reasoning (chains of transitivity steps) as derived
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   958
  concept.  The generic proof elements introduced below depend on
29732
wenzelm
parents: 29729
diff changeset
   959
  rules declared as @{attribute trans} in the context.  It is left to
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   960
  the object-logic to provide a suitable rule collection for mixed
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   961
  relations of \<open>=\<close>, \<open><\<close>, \<open>\<le>\<close>, \<open>\<subset>\<close>,
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   962
  \<open>\<subseteq>\<close> etc.  Due to the flexibility of rule composition
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   963
  (\secref{sec:framework-resolution}), substitution of equals by
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   964
  equals is covered as well, even substitution of inequalities
58552
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 56594
diff changeset
   965
  involving monotonicity conditions; see also @{cite \<open>\S6\<close> "Wenzel-PhD"}
66fed99e874f prefer @{cite} antiquotation;
wenzelm
parents: 56594
diff changeset
   966
  and @{cite "Bauer-Wenzel:2001"}.
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   967
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   968
  The generic calculational mechanism is based on the observation that
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   969
  rules such as \<open>trans:\<close>~@{prop "x = y \<Longrightarrow> y = z \<Longrightarrow> x = z"}
29735
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   970
  proceed from the premises towards the conclusion in a deterministic
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   971
  fashion.  Thus we may reason in forward mode, feeding intermediate
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   972
  results into rules selected from the context.  The course of
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   973
  reasoning is organized by maintaining a secondary fact called
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   974
  ``@{fact calculation}'', apart from the primary ``@{fact this}''
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   975
  already provided by the Isar primitives.  In the definitions below,
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   976
  @{attribute OF} refers to @{inference resolution}
1da96affdefe misc tuning;
wenzelm
parents: 29733
diff changeset
   977
  (\secref{sec:framework-resolution}) with multiple rule arguments,
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   978
  and \<open>trans\<close> represents to a suitable rule from the context:
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   979
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   980
  \begin{matharray}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   981
    @{command "also"}\<open>\<^sub>0\<close> & \equiv & @{command "note"}~\<open>calculation = this\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   982
    @{command "also"}\<open>\<^sub>n\<^sub>+\<^sub>1\<close> & \equiv & @{command "note"}~\<open>calculation = trans [OF calculation this]\<close> \\[0.5ex]
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   983
    @{command "finally"} & \equiv & @{command "also"}~@{command "from"}~\<open>calculation\<close> \\
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   984
  \end{matharray}
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   985
61420
ee42cba50933 redundant due to \parindent 0pt;
wenzelm
parents: 60618
diff changeset
   986
  The start of a calculation is determined implicitly in the
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   987
  text: here @{command also} sets @{fact calculation} to the current
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   988
  result; any subsequent occurrence will update @{fact calculation} by
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   989
  combination with the next result and a transitivity rule.  The
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   990
  calculational sequence is concluded via @{command finally}, where
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   991
  the final result is exposed for use in a concluding claim.
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   992
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   993
  Here is a canonical proof pattern, using @{command have} to
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   994
  establish the intermediate results:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
   995
\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   996
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
   997
(*<*)
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
   998
notepad
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
   999
begin
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
  1000
(*>*)
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
  1001
  have "a = b" sorry
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
  1002
  also have "\<dots> = c" sorry
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
  1003
  also have "\<dots> = d" sorry
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
  1004
  finally have "a = d" .
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
  1005
(*<*)
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40476
diff changeset
  1006
end
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
  1007
(*>*)
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
  1008
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1009
text \<open>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1010
  The term ``\<open>\<dots>\<close>'' above is a special abbreviation
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
  1011
  provided by the Isabelle/Isar syntax layer: it statically refers to
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
  1012
  the right-hand side argument of the previous statement given in the
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
  1013
  text.  Thus it happens to coincide with relevant sub-expressions in
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
  1014
  the calculational chain, but the exact correspondence is dependent
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
  1015
  on the transitivity rules being involved.
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
  1016
61421
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
  1017
  \<^medskip>
e0825405d398 more symbols;
wenzelm
parents: 61420
diff changeset
  1018
  Symmetry rules such as @{prop "x = y \<Longrightarrow> y = x"} are like
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
  1019
  transitivities with only one premise.  Isar maintains a separate
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
  1020
  rule collection declared via the @{attribute sym} attribute, to be
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1021
  used in fact expressions ``\<open>a [symmetric]\<close>'', or single-step
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1022
  proofs ``@{command assume}~\<open>x = y\<close>~@{command then}~@{command
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
  1023
  have}~\<open>y = x\<close>~@{command ".."}''.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58552
diff changeset
  1024
\<close>
29729
c2e926455fcc more on Isar framework -- mostly from Trybulec Festschrift;
wenzelm
parents: 29721
diff changeset
  1025
29742
8edd5198dedb tuned spacing;
wenzelm
parents: 29739
diff changeset
  1026
end