| 
10341
 | 
     1  | 
(* ID:         $Id$ *)
  | 
| 
16417
 | 
     2  | 
theory Blast imports Main begin
  | 
| 
10295
 | 
     3  | 
  | 
| 
 | 
     4  | 
lemma "((\<exists>x. \<forall>y. p(x)=p(y)) = ((\<exists>x. q(x))=(\<forall>y. p(y))))   =    
  | 
| 
 | 
     5  | 
       ((\<exists>x. \<forall>y. q(x)=q(y)) = ((\<exists>x. p(x))=(\<forall>y. q(y))))"
  | 
| 
10844
 | 
     6  | 
by blast
  | 
| 
10295
 | 
     7  | 
  | 
| 
 | 
     8  | 
text{*\noindent Until now, we have proved everything using only induction and
 | 
| 
 | 
     9  | 
simplification.  Substantial proofs require more elaborate types of
  | 
| 
 | 
    10  | 
inference.*}
  | 
| 
 | 
    11  | 
  | 
| 
 | 
    12  | 
lemma "(\<forall>x. honest(x) \<and> industrious(x) \<longrightarrow> healthy(x)) \<and>  
  | 
| 
 | 
    13  | 
       \<not> (\<exists>x. grocer(x) \<and> healthy(x)) \<and> 
  | 
| 
 | 
    14  | 
       (\<forall>x. industrious(x) \<and> grocer(x) \<longrightarrow> honest(x)) \<and> 
  | 
| 
 | 
    15  | 
       (\<forall>x. cyclist(x) \<longrightarrow> industrious(x)) \<and> 
  | 
| 
 | 
    16  | 
       (\<forall>x. \<not>healthy(x) \<and> cyclist(x) \<longrightarrow> \<not>honest(x))  
  | 
| 
 | 
    17  | 
       \<longrightarrow> (\<forall>x. grocer(x) \<longrightarrow> \<not>cyclist(x))";
  | 
| 
10844
 | 
    18  | 
by blast
  | 
| 
10295
 | 
    19  | 
  | 
| 
 | 
    20  | 
lemma "(\<Union>i\<in>I. A(i)) \<inter> (\<Union>j\<in>J. B(j)) =
  | 
| 
 | 
    21  | 
        (\<Union>i\<in>I. \<Union>j\<in>J. A(i) \<inter> B(j))"
  | 
| 
10844
 | 
    22  | 
by blast
  | 
| 
10295
 | 
    23  | 
  | 
| 
 | 
    24  | 
text {*
 | 
| 
 | 
    25  | 
@{thm[display] mult_is_0}
 | 
| 
 | 
    26  | 
 \rulename{mult_is_0}}
 | 
| 
 | 
    27  | 
  | 
| 
 | 
    28  | 
@{thm[display] finite_Un}
 | 
| 
 | 
    29  | 
 \rulename{finite_Un}}
 | 
| 
 | 
    30  | 
*};
  | 
| 
 | 
    31  | 
  | 
| 
 | 
    32  | 
  | 
| 
 | 
    33  | 
lemma [iff]: "(xs@ys = []) = (xs=[] & ys=[])"
  | 
| 
 | 
    34  | 
  apply (induct_tac xs)
  | 
| 
 | 
    35  | 
  by (simp_all);
  | 
| 
 | 
    36  | 
  | 
| 
 | 
    37  | 
(*ideas for uses of intro, etc.: ex/Primes/is_gcd_unique?*)
  | 
| 
 | 
    38  | 
end
  |