doc-src/TutorialI/Rules/Basic.thy
author paulson
Mon, 23 Oct 2000 16:25:04 +0200
changeset 10295 8eb12693cead
child 10341 6eb91805a012
permissions -rw-r--r--
the Rules chapter and theories
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
     1
theory Basic = Main:
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
     2
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
     3
lemma conj_rule: "\<lbrakk> P; Q \<rbrakk> \<Longrightarrow> P \<and> (Q \<and> P)"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
     4
apply (rule conjI)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
     5
 apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
     6
apply (rule conjI)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
     7
 apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
     8
apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
     9
done
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    10
    
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    11
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    12
lemma disj_swap: "P | Q \<Longrightarrow> Q | P"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    13
apply (erule disjE)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    14
 apply (rule disjI2)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    15
 apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    16
apply (rule disjI1)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    17
apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    18
done
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    19
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    20
lemma conj_swap: "P \<and> Q \<Longrightarrow> Q \<and> P"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    21
apply (rule conjI)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    22
 apply (drule conjunct2)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    23
 apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    24
apply (drule conjunct1)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    25
apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    26
done
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    27
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    28
lemma imp_uncurry: "P \<longrightarrow> Q \<longrightarrow> R \<Longrightarrow> P \<and> Q \<longrightarrow> R"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    29
apply (rule impI)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    30
apply (erule conjE)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    31
apply (drule mp)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    32
 apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    33
apply (drule mp)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    34
  apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    35
 apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    36
done
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    37
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    38
text {*
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    39
substitution
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    40
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    41
@{thm[display] ssubst}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    42
\rulename{ssubst}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    43
*};
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    44
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    45
lemma "\<lbrakk> x = f x; P(f x) \<rbrakk> \<Longrightarrow> P x"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    46
apply (erule ssubst)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    47
apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    48
done
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    49
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    50
text {*
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    51
also provable by simp (re-orients)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    52
*};
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    53
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    54
lemma "\<lbrakk> x = f x; P (f x) (f x) x \<rbrakk> \<Longrightarrow> P x x x"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    55
apply (erule ssubst)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    56
back
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    57
back
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    58
back
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    59
back
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    60
apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    61
done
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    62
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    63
text {*
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    64
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ \isadigit{1}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    65
\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    66
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    67
{\isasymlbrakk}x\ {\isacharequal}\ f\ x{\isacharsemicolon}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ x\ x\ x\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    68
\ \isadigit{1}{\isachardot}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    69
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    70
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ \isadigit{1}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    71
\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    72
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    73
{\isasymlbrakk}x\ {\isacharequal}\ f\ x{\isacharsemicolon}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ x\ x\ x\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    74
\ \isadigit{1}{\isachardot}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x\ {\isasymLongrightarrow}\ P\ x\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    75
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    76
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ \isadigit{1}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    77
\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    78
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    79
{\isasymlbrakk}x\ {\isacharequal}\ f\ x{\isacharsemicolon}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ x\ x\ x\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    80
\ \isadigit{1}{\isachardot}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ x\ {\isacharparenleft}f\ x{\isacharparenright}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    81
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    82
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ \isadigit{1}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    83
\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    84
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    85
{\isasymlbrakk}x\ {\isacharequal}\ f\ x{\isacharsemicolon}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ x\ x\ x\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    86
\ \isadigit{1}{\isachardot}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x\ {\isasymLongrightarrow}\ P\ x\ x\ {\isacharparenleft}f\ x{\isacharparenright}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    87
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    88
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ \isadigit{1}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    89
\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    90
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    91
{\isasymlbrakk}x\ {\isacharequal}\ f\ x{\isacharsemicolon}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ x\ x\ x\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    92
\ \isadigit{1}{\isachardot}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    93
*};
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    94
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    95
lemma "\<lbrakk> x = f x; P (f x) (f x) x \<rbrakk> \<Longrightarrow> P x x x"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    96
apply (erule ssubst, assumption)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    97
done
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    98
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    99
lemma "\<lbrakk> x = f x; P (f x) (f x) x \<rbrakk> \<Longrightarrow> P x x x"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   100
apply (erule_tac P="\<lambda>u. P u u x" in ssubst);
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   101
apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   102
done
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   103
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   104
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   105
text {*
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   106
negation
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   107
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   108
@{thm[display] notI}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   109
\rulename{notI}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   110
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   111
@{thm[display] notE}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   112
\rulename{notE}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   113
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   114
@{thm[display] classical}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   115
\rulename{classical}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   116
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   117
@{thm[display] contrapos_pp}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   118
\rulename{contrapos_pp}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   119
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   120
@{thm[display] contrapos_np}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   121
\rulename{contrapos_np}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   122
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   123
@{thm[display] contrapos_nn}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   124
\rulename{contrapos_nn}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   125
*};
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   126
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   127
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   128
lemma "\<lbrakk>\<not>(P\<longrightarrow>Q); \<not>(R\<longrightarrow>Q)\<rbrakk> \<Longrightarrow> R"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   129
apply (erule_tac Q="R\<longrightarrow>Q" in contrapos_np)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   130
txt{*
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   131
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{1}}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   132
\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   133
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   134
{\isasymlbrakk}{\isasymnot}\ {\isacharparenleft}P\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isacharsemicolon}\ {\isasymnot}\ {\isacharparenleft}R\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ R\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   135
\ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}{\isasymnot}\ {\isacharparenleft}P\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isacharsemicolon}\ {\isasymnot}\ R{\isasymrbrakk}\ {\isasymLongrightarrow}\ R\ {\isasymlongrightarrow}\ Q
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   136
*}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   137
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   138
apply intro
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   139
txt{*
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   140
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{3}}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   141
\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   142
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   143
{\isasymlbrakk}{\isasymnot}\ {\isacharparenleft}P\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isacharsemicolon}\ {\isasymnot}\ {\isacharparenleft}R\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ R\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   144
\ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}{\isasymnot}\ {\isacharparenleft}P\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isacharsemicolon}\ {\isasymnot}\ R{\isacharsemicolon}\ R{\isasymrbrakk}\ {\isasymLongrightarrow}\ Q
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   145
*}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   146
apply (erule notE, assumption)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   147
done
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   148
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   149
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   150
lemma "(P \<or> Q) \<and> R \<Longrightarrow> P \<or> Q \<and> R"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   151
apply intro
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   152
txt{*
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   153
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{1}}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   154
\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   155
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   156
{\isacharparenleft}P\ {\isasymor}\ Q{\isacharparenright}\ {\isasymand}\ R\ {\isasymLongrightarrow}\ P\ {\isasymor}\ Q\ {\isasymand}\ R\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   157
\ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}{\isacharparenleft}P\ {\isasymor}\ Q{\isacharparenright}\ {\isasymand}\ R{\isacharsemicolon}\ {\isasymnot}\ {\isacharparenleft}Q\ {\isasymand}\ R{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ P
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   158
*}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   159
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   160
apply (elim conjE disjE)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   161
 apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   162
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   163
txt{*
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   164
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{4}}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   165
\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   166
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   167
{\isacharparenleft}P\ {\isasymor}\ Q{\isacharparenright}\ {\isasymand}\ R\ {\isasymLongrightarrow}\ P\ {\isasymor}\ Q\ {\isasymand}\ R\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   168
\ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}{\isasymnot}\ {\isacharparenleft}Q\ {\isasymand}\ R{\isacharparenright}{\isacharsemicolon}\ R{\isacharsemicolon}\ Q{\isasymrbrakk}\ {\isasymLongrightarrow}\ P
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   169
*}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   170
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   171
apply (erule contrapos_np, rule conjI)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   172
txt{*
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   173
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{6}}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   174
\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   175
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   176
{\isacharparenleft}P\ {\isasymor}\ Q{\isacharparenright}\ {\isasymand}\ R\ {\isasymLongrightarrow}\ P\ {\isasymor}\ Q\ {\isasymand}\ R\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   177
\ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}R{\isacharsemicolon}\ Q{\isacharsemicolon}\ {\isasymnot}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ Q\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   178
\ {\isadigit{2}}{\isachardot}\ {\isasymlbrakk}R{\isacharsemicolon}\ Q{\isacharsemicolon}\ {\isasymnot}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ R
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   179
*}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   180
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   181
  apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   182
 apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   183
done
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   184
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   185
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   186
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   187
text{*Quantifiers*}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   188
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   189
lemma "\<forall>x. P x \<longrightarrow> P x"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   190
apply (rule allI)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   191
apply (rule impI)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   192
apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   193
done
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   194
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   195
lemma "(\<forall>x. P \<longrightarrow> Q x) \<Longrightarrow> P \<longrightarrow> (\<forall>x. Q x)"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   196
apply (rule impI)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   197
apply (rule allI)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   198
apply (drule spec)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   199
apply (drule mp)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   200
  apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   201
 apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   202
done
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   203
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   204
lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (f x); P a\<rbrakk> \<Longrightarrow> P(f (f a))"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   205
apply (frule spec)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   206
apply (drule mp, assumption)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   207
apply (drule spec)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   208
apply (drule mp, assumption, assumption)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   209
done
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   210
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   211
text
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   212
{*
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   213
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{1}}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   214
\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   215
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   216
{\isasymlbrakk}{\isasymforall}x{\isachardot}\ P\ x\ {\isasymlongrightarrow}\ P\ {\isacharparenleft}f\ x{\isacharparenright}{\isacharsemicolon}\ P\ a{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}f\ {\isacharparenleft}f\ a{\isacharparenright}{\isacharparenright}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   217
\ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}{\isasymforall}x{\isachardot}\ P\ x\ {\isasymlongrightarrow}\ P\ {\isacharparenleft}f\ x{\isacharparenright}{\isacharsemicolon}\ P\ a{\isacharsemicolon}\ P\ {\isacharquery}x\ {\isasymlongrightarrow}\ P\ {\isacharparenleft}f\ {\isacharquery}x{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}f\ {\isacharparenleft}f\ a{\isacharparenright}{\isacharparenright}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   218
*}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   219
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   220
text{*
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   221
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{3}}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   222
\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   223
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   224
{\isasymlbrakk}{\isasymforall}x{\isachardot}\ P\ x\ {\isasymlongrightarrow}\ P\ {\isacharparenleft}f\ x{\isacharparenright}{\isacharsemicolon}\ P\ a{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}f\ {\isacharparenleft}f\ a{\isacharparenright}{\isacharparenright}\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   225
\ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}{\isasymforall}x{\isachardot}\ P\ x\ {\isasymlongrightarrow}\ P\ {\isacharparenleft}f\ x{\isacharparenright}{\isacharsemicolon}\ P\ a{\isacharsemicolon}\ P\ {\isacharparenleft}f\ a{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}f\ {\isacharparenleft}f\ a{\isacharparenright}{\isacharparenright}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   226
*}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   227
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   228
lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (f x); P a\<rbrakk> \<Longrightarrow> P(f (f a))"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   229
by blast
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   230
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   231
lemma "(\<exists>x. P x) \<or> (\<exists>x. Q x) \<Longrightarrow> \<exists>x. P x \<or> Q x"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   232
apply elim
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   233
 apply intro
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   234
 apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   235
apply (intro exI disjI2) (*or else we get disjCI*)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   236
apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   237
done
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   238
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   239
lemma "(P\<longrightarrow>Q) \<or> (Q\<longrightarrow>P)"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   240
apply intro
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   241
apply elim
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   242
apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   243
done
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   244
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   245
lemma "(P\<or>Q)\<and>(P\<or>R) \<Longrightarrow> P \<or> (Q\<and>R)"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   246
apply intro
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   247
apply (elim conjE disjE)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   248
apply blast
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   249
apply blast
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   250
apply blast
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   251
apply blast
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   252
(*apply elim*)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   253
done
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   254
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   255
lemma "(\<exists>x. P \<and> Q x) \<Longrightarrow> P \<and> (\<exists>x. Q x)"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   256
apply (erule exE)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   257
apply (erule conjE)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   258
apply (rule conjI)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   259
 apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   260
apply (rule exI)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   261
 apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   262
done
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   263
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   264
lemma "(\<exists>x. P x) \<and> (\<exists>x. Q x) \<Longrightarrow> \<exists>x. P x \<and> Q x"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   265
apply (erule conjE)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   266
apply (erule exE)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   267
apply (erule exE)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   268
apply (rule exI)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   269
apply (rule conjI)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   270
 apply assumption
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   271
oops
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   272
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   273
lemma "\<forall> z. R z z \<Longrightarrow> \<exists>x. \<forall> y. R x y"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   274
apply (rule exI)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   275
apply (rule allI)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   276
apply (drule spec)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   277
oops
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   278
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   279
lemma "\<forall>x. \<exists> y. x=y"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   280
apply (rule allI)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   281
apply (rule exI)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   282
apply (rule refl)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   283
done
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   284
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   285
lemma "\<exists>x. \<forall> y. x=y"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   286
apply (rule exI)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   287
apply (rule allI)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   288
oops
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   289
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   290
end