src/HOL/Isar_Examples/Basic_Logic.thy
author wenzelm
Sat, 10 Oct 2020 20:31:54 +0200
changeset 72421 9a8bc089890d
parent 63585 f4a308fdf664
child 76987 4c275405faae
permissions -rw-r--r--
support for MSYS2/MinGW64 on Windows;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33026
8f35633c4922 modernized session Isar_Examples;
wenzelm
parents: 31758
diff changeset
     1
(*  Title:      HOL/Isar_Examples/Basic_Logic.thy
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
     2
    Author:     Makarius
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
     3
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
     4
Basic propositional and quantifier reasoning.
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
     5
*)
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
     6
58882
6e2010ab8bd9 modernized header;
wenzelm
parents: 58614
diff changeset
     7
section \<open>Basic logical reasoning\<close>
7748
5b9c45b21782 improved presentation;
wenzelm
parents: 7740
diff changeset
     8
31758
3edd5f813f01 observe standard theory naming conventions;
wenzelm
parents: 23393
diff changeset
     9
theory Basic_Logic
63585
wenzelm
parents: 61932
diff changeset
    10
  imports Main
31758
3edd5f813f01 observe standard theory naming conventions;
wenzelm
parents: 23393
diff changeset
    11
begin
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    12
7761
7fab9592384f improved presentation;
wenzelm
parents: 7748
diff changeset
    13
58614
7338eb25226c more cartouches;
wenzelm
parents: 55656
diff changeset
    14
subsection \<open>Pure backward reasoning\<close>
7740
2fbe5ce9845f tuned comments;
wenzelm
parents: 7604
diff changeset
    15
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    16
text \<open>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    17
  In order to get a first idea of how Isabelle/Isar proof documents may look
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    18
  like, we consider the propositions \<open>I\<close>, \<open>K\<close>, and \<open>S\<close>. The following (rather
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    19
  explicit) proofs should require little extra explanations.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    20
\<close>
7001
8121e11ed765 Deriving rules in Isabelle;
wenzelm
parents: 6892
diff changeset
    21
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    22
lemma I: "A \<longrightarrow> A"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    23
proof
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    24
  assume A
23393
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23373
diff changeset
    25
  show A by fact
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    26
qed
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    27
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    28
lemma K: "A \<longrightarrow> B \<longrightarrow> A"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    29
proof
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    30
  assume A
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    31
  show "B \<longrightarrow> A"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    32
  proof
23393
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23373
diff changeset
    33
    show A by fact
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    34
  qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    35
qed
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    36
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    37
lemma S: "(A \<longrightarrow> B \<longrightarrow> C) \<longrightarrow> (A \<longrightarrow> B) \<longrightarrow> A \<longrightarrow> C"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    38
proof
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    39
  assume "A \<longrightarrow> B \<longrightarrow> C"
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    40
  show "(A \<longrightarrow> B) \<longrightarrow> A \<longrightarrow> C"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    41
  proof
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    42
    assume "A \<longrightarrow> B"
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    43
    show "A \<longrightarrow> C"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    44
    proof
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    45
      assume A
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    46
      show C
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    47
      proof (rule mp)
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    48
        show "B \<longrightarrow> C" by (rule mp) fact+
23393
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23373
diff changeset
    49
        show B by (rule mp) fact+
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    50
      qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    51
    qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    52
  qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    53
qed
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    54
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    55
text \<open>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    56
  Isar provides several ways to fine-tune the reasoning, avoiding excessive
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    57
  detail. Several abbreviated language elements are available, enabling the
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    58
  writer to express proofs in a more concise way, even without referring to
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    59
  any automated proof tools yet.
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
    60
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    61
  Concluding any (sub-)proof already involves solving any remaining goals by
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    62
  assumption\<^footnote>\<open>This is not a completely trivial operation, as proof by
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    63
  assumption may involve full higher-order unification.\<close>. Thus we may skip the
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    64
  rather vacuous body of the above proof.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    65
\<close>
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
    66
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    67
lemma "A \<longrightarrow> A"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    68
proof
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    69
qed
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
    70
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    71
text \<open>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    72
  Note that the \<^theory_text>\<open>proof\<close> command refers to the \<open>rule\<close> method (without
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    73
  arguments) by default. Thus it implicitly applies a single rule, as
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    74
  determined from the syntactic form of the statements involved. The \<^theory_text>\<open>by\<close>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    75
  command abbreviates any proof with empty body, so the proof may be further
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    76
  pruned.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    77
\<close>
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
    78
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    79
lemma "A \<longrightarrow> A"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    80
  by rule
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
    81
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    82
text \<open>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    83
  Proof by a single rule may be abbreviated as double-dot.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    84
\<close>
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
    85
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    86
lemma "A \<longrightarrow> A" ..
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
    87
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    88
text \<open>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    89
  Thus we have arrived at an adequate representation of the proof of a
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    90
  tautology that holds by a single standard rule.\<^footnote>\<open>Apparently, the
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    91
  rule here is implication introduction.\<close>
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
    92
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    93
  \<^medskip>
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    94
  Let us also reconsider \<open>K\<close>. Its statement is composed of iterated
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    95
  connectives. Basic decomposition is by a single rule at a time, which is why
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    96
  our first version above was by nesting two proofs.
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
    97
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    98
  The \<open>intro\<close> proof method repeatedly decomposes a goal's conclusion.\<^footnote>\<open>The
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
    99
  dual method is \<open>elim\<close>, acting on a goal's premises.\<close>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   100
\<close>
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
   101
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   102
lemma "A \<longrightarrow> B \<longrightarrow> A"
12387
fe2353a8d1e8 fixed intro steps;
wenzelm
parents: 10636
diff changeset
   103
proof (intro impI)
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   104
  assume A
23393
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23373
diff changeset
   105
  show A by fact
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   106
qed
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
   107
58614
7338eb25226c more cartouches;
wenzelm
parents: 55656
diff changeset
   108
text \<open>Again, the body may be collapsed.\<close>
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
   109
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   110
lemma "A \<longrightarrow> B \<longrightarrow> A"
12387
fe2353a8d1e8 fixed intro steps;
wenzelm
parents: 10636
diff changeset
   111
  by (intro impI)
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
   112
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   113
text \<open>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   114
  Just like \<open>rule\<close>, the \<open>intro\<close> and \<open>elim\<close> proof methods pick standard
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   115
  structural rules, in case no explicit arguments are given. While implicit
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   116
  rules are usually just fine for single rule application, this may go too far
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   117
  with iteration. Thus in practice, \<open>intro\<close> and \<open>elim\<close> would be typically
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   118
  restricted to certain structures by giving a few rules only, e.g.\ \<^theory_text>\<open>proof
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   119
  (intro impI allI)\<close> to strip implications and universal quantifiers.
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
   120
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   121
  Such well-tuned iterated decomposition of certain structures is the prime
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   122
  application of \<open>intro\<close> and \<open>elim\<close>. In contrast, terminal steps that solve a
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   123
  goal completely are usually performed by actual automated proof methods
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   124
  (such as \<^theory_text>\<open>by blast\<close>.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   125
\<close>
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
   126
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
   127
58614
7338eb25226c more cartouches;
wenzelm
parents: 55656
diff changeset
   128
subsection \<open>Variations of backward vs.\ forward reasoning\<close>
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
   129
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   130
text \<open>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   131
  Certainly, any proof may be performed in backward-style only. On the other
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   132
  hand, small steps of reasoning are often more naturally expressed in
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   133
  forward-style. Isar supports both backward and forward reasoning as a
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   134
  first-class concept. In order to demonstrate the difference, we consider
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   135
  several proofs of \<open>A \<and> B \<longrightarrow> B \<and> A\<close>.
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
   136
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   137
  The first version is purely backward.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   138
\<close>
7001
8121e11ed765 Deriving rules in Isabelle;
wenzelm
parents: 6892
diff changeset
   139
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   140
lemma "A \<and> B \<longrightarrow> B \<and> A"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   141
proof
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   142
  assume "A \<and> B"
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   143
  show "B \<and> A"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   144
  proof
23393
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23373
diff changeset
   145
    show B by (rule conjunct2) fact
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23373
diff changeset
   146
    show A by (rule conjunct1) fact
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   147
  qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   148
qed
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   149
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   150
text \<open>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   151
  Above, the projection rules \<open>conjunct1\<close> / \<open>conjunct2\<close> had to be named
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   152
  explicitly, since the goals \<open>B\<close> and \<open>A\<close> did not provide any structural clue.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   153
  This may be avoided using \<^theory_text>\<open>from\<close> to focus on the \<open>A \<and> B\<close> assumption as the
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   154
  current facts, enabling the use of double-dot proofs. Note that \<^theory_text>\<open>from\<close>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   155
  already does forward-chaining, involving the \<open>conjE\<close> rule here.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   156
\<close>
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   157
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   158
lemma "A \<and> B \<longrightarrow> B \<and> A"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   159
proof
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   160
  assume "A \<and> B"
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   161
  show "B \<and> A"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   162
  proof
58614
7338eb25226c more cartouches;
wenzelm
parents: 55656
diff changeset
   163
    from \<open>A \<and> B\<close> show B ..
7338eb25226c more cartouches;
wenzelm
parents: 55656
diff changeset
   164
    from \<open>A \<and> B\<close> show A ..
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   165
  qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   166
qed
7604
wenzelm
parents: 7480
diff changeset
   167
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   168
text \<open>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   169
  In the next version, we move the forward step one level upwards.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   170
  Forward-chaining from the most recent facts is indicated by the \<^theory_text>\<open>then\<close>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   171
  command. Thus the proof of \<open>B \<and> A\<close> from \<open>A \<and> B\<close> actually becomes an
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   172
  elimination, rather than an introduction. The resulting proof structure
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   173
  directly corresponds to that of the \<open>conjE\<close> rule, including the repeated
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   174
  goal proposition that is abbreviated as \<open>?thesis\<close> below.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   175
\<close>
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
   176
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   177
lemma "A \<and> B \<longrightarrow> B \<and> A"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   178
proof
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   179
  assume "A \<and> B"
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   180
  then show "B \<and> A"
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61542
diff changeset
   181
  proof                    \<comment> \<open>rule \<open>conjE\<close> of \<open>A \<and> B\<close>\<close>
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 18193
diff changeset
   182
    assume B A
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61542
diff changeset
   183
    then show ?thesis ..   \<comment> \<open>rule \<open>conjI\<close> of \<open>B \<and> A\<close>\<close>
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   184
  qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   185
qed
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
   186
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   187
text \<open>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   188
  In the subsequent version we flatten the structure of the main body by doing
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   189
  forward reasoning all the time. Only the outermost decomposition step is
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   190
  left as backward.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   191
\<close>
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
   192
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   193
lemma "A \<and> B \<longrightarrow> B \<and> A"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   194
proof
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   195
  assume "A \<and> B"
58614
7338eb25226c more cartouches;
wenzelm
parents: 55656
diff changeset
   196
  from \<open>A \<and> B\<close> have A ..
7338eb25226c more cartouches;
wenzelm
parents: 55656
diff changeset
   197
  from \<open>A \<and> B\<close> have B ..
7338eb25226c more cartouches;
wenzelm
parents: 55656
diff changeset
   198
  from \<open>B\<close> \<open>A\<close> show "B \<and> A" ..
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   199
qed
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   200
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   201
text \<open>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   202
  We can still push forward-reasoning a bit further, even at the risk of
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   203
  getting ridiculous. Note that we force the initial proof step to do nothing
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   204
  here, by referring to the \<open>-\<close> proof method.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   205
\<close>
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
   206
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   207
lemma "A \<and> B \<longrightarrow> B \<and> A"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   208
proof -
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   209
  {
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   210
    assume "A \<and> B"
58614
7338eb25226c more cartouches;
wenzelm
parents: 55656
diff changeset
   211
    from \<open>A \<and> B\<close> have A ..
7338eb25226c more cartouches;
wenzelm
parents: 55656
diff changeset
   212
    from \<open>A \<and> B\<close> have B ..
7338eb25226c more cartouches;
wenzelm
parents: 55656
diff changeset
   213
    from \<open>B\<close> \<open>A\<close> have "B \<and> A" ..
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   214
  }
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61542
diff changeset
   215
  then show ?thesis ..         \<comment> \<open>rule \<open>impI\<close>\<close>
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   216
qed
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
   217
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   218
text \<open>
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   219
  \<^medskip>
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   220
  With these examples we have shifted through a whole range from purely
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   221
  backward to purely forward reasoning. Apparently, in the extreme ends we get
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   222
  slightly ill-structured proofs, which also require much explicit naming of
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   223
  either rules (backward) or local facts (forward).
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
   224
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   225
  The general lesson learned here is that good proof style would achieve just
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   226
  the \<^emph>\<open>right\<close> balance of top-down backward decomposition, and bottom-up
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   227
  forward composition. In general, there is no single best way to arrange some
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   228
  pieces of formal reasoning, of course. Depending on the actual applications,
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   229
  the intended audience etc., rules (and methods) on the one hand vs.\ facts
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   230
  on the other hand have to be emphasized in an appropriate way. This requires
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   231
  the proof writer to develop good taste, and some practice, of course.
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
   232
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   233
  \<^medskip>
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   234
  For our example the most appropriate way of reasoning is probably the middle
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   235
  one, with conjunction introduction done after elimination.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   236
\<close>
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
   237
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   238
lemma "A \<and> B \<longrightarrow> B \<and> A"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   239
proof
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   240
  assume "A \<and> B"
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   241
  then show "B \<and> A"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   242
  proof
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 18193
diff changeset
   243
    assume B A
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 18193
diff changeset
   244
    then show ?thesis ..
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   245
  qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   246
qed
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
   247
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
   248
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   249
58614
7338eb25226c more cartouches;
wenzelm
parents: 55656
diff changeset
   250
subsection \<open>A few examples from ``Introduction to Isabelle''\<close>
7001
8121e11ed765 Deriving rules in Isabelle;
wenzelm
parents: 6892
diff changeset
   251
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   252
text \<open>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   253
  We rephrase some of the basic reasoning examples of @{cite
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   254
  "isabelle-intro"}, using HOL rather than FOL.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   255
\<close>
37671
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
   256
7820
cad7cc30fa40 more explanations;
wenzelm
parents: 7761
diff changeset
   257
58614
7338eb25226c more cartouches;
wenzelm
parents: 55656
diff changeset
   258
subsubsection \<open>A propositional proof\<close>
7833
f5288e4b95d1 improved presentation;
wenzelm
parents: 7820
diff changeset
   259
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   260
text \<open>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   261
  We consider the proposition \<open>P \<or> P \<longrightarrow> P\<close>. The proof below involves
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   262
  forward-chaining from \<open>P \<or> P\<close>, followed by an explicit case-analysis on the
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   263
  two \<^emph>\<open>identical\<close> cases.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   264
\<close>
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   265
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   266
lemma "P \<or> P \<longrightarrow> P"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   267
proof
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   268
  assume "P \<or> P"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 18193
diff changeset
   269
  then show P
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61542
diff changeset
   270
  proof                    \<comment> \<open>rule \<open>disjE\<close>: \smash{$\infer{\<open>C\<close>}{\<open>A \<or> B\<close> & \infer*{\<open>C\<close>}{[\<open>A\<close>]} & \infer*{\<open>C\<close>}{[\<open>B\<close>]}}$}\<close>
23393
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23373
diff changeset
   271
    assume P show P by fact
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   272
  next
23393
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23373
diff changeset
   273
    assume P show P by fact
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   274
  qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   275
qed
7833
f5288e4b95d1 improved presentation;
wenzelm
parents: 7820
diff changeset
   276
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   277
text \<open>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   278
  Case splits are \<^emph>\<open>not\<close> hardwired into the Isar language as a special
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   279
  feature. The \<^theory_text>\<open>next\<close> command used to separate the cases above is just a
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   280
  short form of managing block structure.
7833
f5288e4b95d1 improved presentation;
wenzelm
parents: 7820
diff changeset
   281
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   282
  \<^medskip>
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   283
  In general, applying proof methods may split up a goal into separate
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   284
  ``cases'', i.e.\ new subgoals with individual local assumptions. The
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   285
  corresponding proof text typically mimics this by establishing results in
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   286
  appropriate contexts, separated by blocks.
7833
f5288e4b95d1 improved presentation;
wenzelm
parents: 7820
diff changeset
   287
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   288
  In order to avoid too much explicit parentheses, the Isar system implicitly
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   289
  opens an additional block for any new goal, the \<^theory_text>\<open>next\<close> statement then
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   290
  closes one block level, opening a new one. The resulting behaviour is what
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   291
  one would expect from separating cases, only that it is more flexible. E.g.\
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   292
  an induction base case (which does not introduce local assumptions) would
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   293
  \<^emph>\<open>not\<close> require \<^theory_text>\<open>next\<close> to separate the subsequent step case.
7833
f5288e4b95d1 improved presentation;
wenzelm
parents: 7820
diff changeset
   294
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   295
  \<^medskip>
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   296
  In our example the situation is even simpler, since the two cases actually
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   297
  coincide. Consequently the proof may be rephrased as follows.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   298
\<close>
7833
f5288e4b95d1 improved presentation;
wenzelm
parents: 7820
diff changeset
   299
55656
eb07b0acbebc more symbols;
wenzelm
parents: 55640
diff changeset
   300
lemma "P \<or> P \<longrightarrow> P"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   301
proof
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   302
  assume "P \<or> P"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 18193
diff changeset
   303
  then show P
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   304
  proof
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   305
    assume P
23393
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23373
diff changeset
   306
    show P by fact
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23373
diff changeset
   307
    show P by fact
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   308
  qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   309
qed
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   310
58614
7338eb25226c more cartouches;
wenzelm
parents: 55656
diff changeset
   311
text \<open>Again, the rather vacuous body of the proof may be collapsed.
37671
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
   312
  Thus the case analysis degenerates into two assumption steps, which
fa53d267dab3 misc tuning and modernization;
wenzelm
parents: 33026
diff changeset
   313
  are implicitly performed when concluding the single rule step of the
58614
7338eb25226c more cartouches;
wenzelm
parents: 55656
diff changeset
   314
  double-dot proof as follows.\<close>
7833
f5288e4b95d1 improved presentation;
wenzelm
parents: 7820
diff changeset
   315
55656
eb07b0acbebc more symbols;
wenzelm
parents: 55640
diff changeset
   316
lemma "P \<or> P \<longrightarrow> P"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   317
proof
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   318
  assume "P \<or> P"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 18193
diff changeset
   319
  then show P ..
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   320
qed
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   321
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   322
58614
7338eb25226c more cartouches;
wenzelm
parents: 55656
diff changeset
   323
subsubsection \<open>A quantifier proof\<close>
7833
f5288e4b95d1 improved presentation;
wenzelm
parents: 7820
diff changeset
   324
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   325
text \<open>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   326
  To illustrate quantifier reasoning, let us prove
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   327
  \<open>(\<exists>x. P (f x)) \<longrightarrow> (\<exists>y. P y)\<close>. Informally, this holds because any \<open>a\<close> with
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   328
  \<open>P (f a)\<close> may be taken as a witness for the second existential statement.
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   329
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   330
  The first proof is rather verbose, exhibiting quite a lot of (redundant)
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   331
  detail. It gives explicit rules, even with some instantiation. Furthermore,
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   332
  we encounter two new language elements: the \<^theory_text>\<open>fix\<close> command augments the
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   333
  context by some new ``arbitrary, but fixed'' element; the \<^theory_text>\<open>is\<close> annotation
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   334
  binds term abbreviations by higher-order pattern matching.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   335
\<close>
7833
f5288e4b95d1 improved presentation;
wenzelm
parents: 7820
diff changeset
   336
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   337
lemma "(\<exists>x. P (f x)) \<longrightarrow> (\<exists>y. P y)"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   338
proof
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   339
  assume "\<exists>x. P (f x)"
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   340
  then show "\<exists>y. P y"
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61542
diff changeset
   341
  proof (rule exE)             \<comment> \<open>rule \<open>exE\<close>: \smash{$\infer{\<open>B\<close>}{\<open>\<exists>x. A(x)\<close> & \infer*{\<open>B\<close>}{[\<open>A(x)\<close>]_x}}$}\<close>
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   342
    fix a
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   343
    assume "P (f a)" (is "P ?witness")
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 18193
diff changeset
   344
    then show ?thesis by (rule exI [of P ?witness])
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   345
  qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   346
qed
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   347
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   348
text \<open>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   349
  While explicit rule instantiation may occasionally improve readability of
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   350
  certain aspects of reasoning, it is usually quite redundant. Above, the
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   351
  basic proof outline gives already enough structural clues for the system to
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   352
  infer both the rules and their instances (by higher-order unification). Thus
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   353
  we may as well prune the text as follows.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   354
\<close>
7833
f5288e4b95d1 improved presentation;
wenzelm
parents: 7820
diff changeset
   355
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   356
lemma "(\<exists>x. P (f x)) \<longrightarrow> (\<exists>y. P y)"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   357
proof
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   358
  assume "\<exists>x. P (f x)"
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   359
  then show "\<exists>y. P y"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   360
  proof
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   361
    fix a
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   362
    assume "P (f a)"
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 18193
diff changeset
   363
    then show ?thesis ..
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   364
  qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   365
qed
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   366
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   367
text \<open>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   368
  Explicit \<open>\<exists>\<close>-elimination as seen above can become quite cumbersome in
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   369
  practice. The derived Isar language element ``\<^theory_text>\<open>obtain\<close>'' provides a more
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   370
  handsome way to do generalized existence reasoning.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   371
\<close>
9477
9506127f6fbb obtain;
wenzelm
parents: 8902
diff changeset
   372
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   373
lemma "(\<exists>x. P (f x)) \<longrightarrow> (\<exists>y. P y)"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   374
proof
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   375
  assume "\<exists>x. P (f x)"
10636
wenzelm
parents: 10007
diff changeset
   376
  then obtain a where "P (f a)" ..
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   377
  then show "\<exists>y. P y" ..
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   378
qed
9477
9506127f6fbb obtain;
wenzelm
parents: 8902
diff changeset
   379
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   380
text \<open>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   381
  Technically, \<^theory_text>\<open>obtain\<close> is similar to \<^theory_text>\<open>fix\<close> and \<^theory_text>\<open>assume\<close> together with a
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   382
  soundness proof of the elimination involved. Thus it behaves similar to any
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   383
  other forward proof element. Also note that due to the nature of general
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   384
  existence reasoning involved here, any result exported from the context of
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   385
  an \<^theory_text>\<open>obtain\<close> statement may \<^emph>\<open>not\<close> refer to the parameters introduced there.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   386
\<close>
9477
9506127f6fbb obtain;
wenzelm
parents: 8902
diff changeset
   387
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   388
58614
7338eb25226c more cartouches;
wenzelm
parents: 55656
diff changeset
   389
subsubsection \<open>Deriving rules in Isabelle\<close>
7001
8121e11ed765 Deriving rules in Isabelle;
wenzelm
parents: 6892
diff changeset
   390
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   391
text \<open>
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   392
  We derive the conjunction elimination rule from the corresponding
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   393
  projections. The proof is quite straight-forward, since Isabelle/Isar
61932
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   394
  supports non-atomic goals and assumptions fully transparently.
2e48182cc82c misc tuning and modernization;
wenzelm
parents: 61799
diff changeset
   395
\<close>
7001
8121e11ed765 Deriving rules in Isabelle;
wenzelm
parents: 6892
diff changeset
   396
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   397
theorem conjE: "A \<and> B \<Longrightarrow> (A \<Longrightarrow> B \<Longrightarrow> C) \<Longrightarrow> C"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   398
proof -
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   399
  assume "A \<and> B"
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
   400
  assume r: "A \<Longrightarrow> B \<Longrightarrow> C"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   401
  show C
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   402
  proof (rule r)
23393
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23373
diff changeset
   403
    show A by (rule conjunct1) fact
31781b2de73d tuned proofs: avoid implicit prems;
wenzelm
parents: 23373
diff changeset
   404
    show B by (rule conjunct2) fact
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   405
  qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   406
qed
7001
8121e11ed765 Deriving rules in Isabelle;
wenzelm
parents: 6892
diff changeset
   407
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   408
end