author | wenzelm |
Sat, 18 May 2019 12:08:30 +0200 | |
changeset 70281 | 110df6f91376 |
parent 67406 | 23307fd33906 |
permissions | -rw-r--r-- |
16417 | 1 |
theory Force imports Main begin |
11456 | 2 |
(*Use Divides rather than Main to experiment with the first proof. |
3 |
Otherwise, it is done by the nat_divide_cancel_factor simprocs.*) |
|
11407 | 4 |
|
58860 | 5 |
declare div_mult_self_is_m [simp del] |
10295 | 6 |
|
11407 | 7 |
lemma div_mult_self_is_m: |
8 |
"0<n \<Longrightarrow> (m*n) div n = (m::nat)" |
|
64242
93c6f0da5c70
more standardized theorem names for facts involving the div and mod identity
haftmann
parents:
58860
diff
changeset
|
9 |
apply (insert div_mult_mod_eq [of "m*n" n]) |
11407 | 10 |
apply simp |
11 |
done |
|
10295 | 12 |
|
13 |
||
14 |
lemma "(\<forall>x. P x) \<and> (\<exists>x. Q x) \<longrightarrow> (\<forall>x. P x \<and> Q x)" |
|
15 |
apply clarify |
|
16 |
oops |
|
17 |
||
67406 | 18 |
text \<open> |
10295 | 19 |
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{1}}\isanewline |
20 |
\isanewline |
|
21 |
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline |
|
22 |
{\isacharparenleft}{\isasymforall}x{\isachardot}\ P\ x{\isacharparenright}\ {\isasymand}\ {\isacharparenleft}{\isasymexists}x{\isachardot}\ Q\ x{\isacharparenright}\ {\isasymlongrightarrow}\ {\isacharparenleft}{\isasymforall}x{\isachardot}\ P\ x\ {\isasymand}\ Q\ x{\isacharparenright}\isanewline |
|
23 |
\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ xa{\isachardot}\ {\isasymlbrakk}{\isasymforall}x{\isachardot}\ P\ x{\isacharsemicolon}\ Q\ xa{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ x\ {\isasymand}\ Q\ x |
|
67406 | 24 |
\<close> |
10295 | 25 |
|
67406 | 26 |
text \<open> |
10295 | 27 |
couldn't find a good example of clarsimp |
28 |
||
29 |
@{thm[display]"someI"} |
|
30 |
\rulename{someI} |
|
67406 | 31 |
\<close> |
10295 | 32 |
|
33 |
lemma "\<lbrakk>Q a; P a\<rbrakk> \<Longrightarrow> P (SOME x. P x \<and> Q x) \<and> Q (SOME x. P x \<and> Q x)" |
|
34 |
apply (rule someI) |
|
35 |
oops |
|
36 |
||
37 |
lemma "\<lbrakk>Q a; P a\<rbrakk> \<Longrightarrow> P (SOME x. P x \<and> Q x) \<and> Q (SOME x. P x \<and> Q x)" |
|
38 |
apply (fast intro!: someI) |
|
39 |
done |
|
40 |
||
67406 | 41 |
text\<open> |
10295 | 42 |
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ \isadigit{1}\isanewline |
43 |
\isanewline |
|
44 |
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline |
|
45 |
{\isasymlbrakk}Q\ a{\isacharsemicolon}\ P\ a{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}SOME\ x{\isachardot}\ P\ x\ {\isasymand}\ Q\ x{\isacharparenright}\ {\isasymand}\ Q\ {\isacharparenleft}SOME\ x{\isachardot}\ P\ x\ {\isasymand}\ Q\ x{\isacharparenright}\isanewline |
|
46 |
\ \isadigit{1}{\isachardot}\ {\isasymlbrakk}Q\ a{\isacharsemicolon}\ P\ a{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isacharquery}x\ {\isasymand}\ Q\ {\isacharquery}x |
|
67406 | 47 |
\<close> |
10295 | 48 |
|
49 |
end |
|
50 |