equal
deleted
inserted
replaced
1 theory Force imports Main begin |
1 theory Force imports Main begin |
2 (*Use Divides rather than Main to experiment with the first proof. |
2 (*Use Divides rather than Main to experiment with the first proof. |
3 Otherwise, it is done by the nat_divide_cancel_factor simprocs.*) |
3 Otherwise, it is done by the nat_divide_cancel_factor simprocs.*) |
4 |
4 |
5 declare div_mult_self_is_m [simp del]; |
5 declare div_mult_self_is_m [simp del] |
6 |
6 |
7 lemma div_mult_self_is_m: |
7 lemma div_mult_self_is_m: |
8 "0<n \<Longrightarrow> (m*n) div n = (m::nat)" |
8 "0<n \<Longrightarrow> (m*n) div n = (m::nat)" |
9 apply (insert mod_div_equality [of "m*n" n]) |
9 apply (insert mod_div_equality [of "m*n" n]) |
10 apply simp |
10 apply simp |
19 proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{1}}\isanewline |
19 proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{1}}\isanewline |
20 \isanewline |
20 \isanewline |
21 goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\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 |
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 |
23 \ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ xa{\isachardot}\ {\isasymlbrakk}{\isasymforall}x{\isachardot}\ P\ x{\isacharsemicolon}\ Q\ xa{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ x\ {\isasymand}\ Q\ x |
24 *}; |
24 *} |
25 |
25 |
26 text {* |
26 text {* |
27 couldn't find a good example of clarsimp |
27 couldn't find a good example of clarsimp |
28 |
28 |
29 @{thm[display]"someI"} |
29 @{thm[display]"someI"} |
30 \rulename{someI} |
30 \rulename{someI} |
31 *}; |
31 *} |
32 |
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)" |
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) |
34 apply (rule someI) |
35 oops |
35 oops |
36 |
36 |