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