| 
10341
 | 
     1  | 
(* ID:         $Id$ *)
  | 
| 
16417
 | 
     2  | 
theory Force imports Main begin 
  | 
| 
11456
 | 
     3  | 
  (*Use Divides rather than Main to experiment with the first proof.
  | 
| 
 | 
     4  | 
    Otherwise, it is done by the nat_divide_cancel_factor simprocs.*)
  | 
| 
11407
 | 
     5  | 
  | 
| 
 | 
     6  | 
declare div_mult_self_is_m [simp del];
  | 
| 
10295
 | 
     7  | 
  | 
| 
11407
 | 
     8  | 
lemma div_mult_self_is_m: 
  | 
| 
 | 
     9  | 
      "0<n \<Longrightarrow> (m*n) div n = (m::nat)"
  | 
| 
 | 
    10  | 
apply (insert mod_div_equality [of "m*n" n])
  | 
| 
 | 
    11  | 
apply simp
  | 
| 
 | 
    12  | 
done
  | 
| 
10295
 | 
    13  | 
  | 
| 
 | 
    14  | 
  | 
| 
 | 
    15  | 
lemma "(\<forall>x. P x) \<and> (\<exists>x. Q x) \<longrightarrow> (\<forall>x. P x \<and> Q x)"
  | 
| 
 | 
    16  | 
apply clarify
  | 
| 
 | 
    17  | 
oops
  | 
| 
 | 
    18  | 
  | 
| 
 | 
    19  | 
text {*
 | 
| 
 | 
    20  | 
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{1}}\isanewline
 | 
| 
 | 
    21  | 
\isanewline
  | 
| 
 | 
    22  | 
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
 | 
| 
 | 
    23  | 
{\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
 | 
| 
 | 
    24  | 
\ {\isadigit{1}}{\isachardot}\ {\isasymAnd}x\ xa{\isachardot}\ {\isasymlbrakk}{\isasymforall}x{\isachardot}\ P\ x{\isacharsemicolon}\ Q\ xa{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ x\ {\isasymand}\ Q\ x
 | 
| 
 | 
    25  | 
*};
  | 
| 
 | 
    26  | 
  | 
| 
 | 
    27  | 
text {*
 | 
| 
 | 
    28  | 
couldn't find a good example of clarsimp
  | 
| 
 | 
    29  | 
  | 
| 
 | 
    30  | 
@{thm[display]"someI"}
 | 
| 
 | 
    31  | 
\rulename{someI}
 | 
| 
 | 
    32  | 
*};
  | 
| 
 | 
    33  | 
  | 
| 
 | 
    34  | 
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)"
  | 
| 
 | 
    35  | 
apply (rule someI)
  | 
| 
 | 
    36  | 
oops
  | 
| 
 | 
    37  | 
  | 
| 
 | 
    38  | 
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)"
  | 
| 
 | 
    39  | 
apply (fast intro!: someI)
  | 
| 
 | 
    40  | 
done
  | 
| 
 | 
    41  | 
  | 
| 
 | 
    42  | 
text{*
 | 
| 
 | 
    43  | 
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ \isadigit{1}\isanewline
 | 
| 
 | 
    44  | 
\isanewline
  | 
| 
 | 
    45  | 
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
 | 
| 
 | 
    46  | 
{\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
 | 
| 
 | 
    47  | 
\ \isadigit{1}{\isachardot}\ {\isasymlbrakk}Q\ a{\isacharsemicolon}\ P\ a{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isacharquery}x\ {\isasymand}\ Q\ {\isacharquery}x
 | 
| 
 | 
    48  | 
*}
  | 
| 
 | 
    49  | 
  | 
| 
 | 
    50  | 
end
  | 
| 
 | 
    51  | 
  |