Theory Drinker

Up to index of Isabelle/HOL/Isar_examples

theory Drinker
imports Main
begin

(*  Title:      HOL/Isar_examples/Drinker.thy
    ID:         $Id: Drinker.thy,v 1.3 2005/06/22 17:41:16 wenzelm Exp $
    Author:     Makarius
*)

header {* The Drinker's Principle *}

theory Drinker imports Main begin

text {*
  Here is another example of classical reasoning: the Drinker's
  Principle says that for some person, if he is drunk, everybody else
  is drunk!

  We first prove a classical part of de-Morgan's law.
*}

lemma deMorgan:
  assumes "¬ (∀x. P x)"
  shows "∃x. ¬ P x"
  using prems
proof (rule contrapos_np)
  assume a: "¬ (∃x. ¬ P x)"
  show "∀x. P x"
  proof
    fix x
    show "P x"
    proof (rule classical)
      assume "¬ P x"
      then have "∃x. ¬ P x" ..
      with a show ?thesis by contradiction
    qed
  qed
qed

theorem Drinker's_Principle: "∃x. drunk x --> (∀x. drunk x)"
proof cases
  fix a assume "∀x. drunk x"
  then have "drunk a --> (∀x. drunk x)" ..
  then show ?thesis ..
next
  assume "¬ (∀x. drunk x)"
  then have "∃x. ¬ drunk x" by (rule deMorgan)
  then obtain a where a: "¬ drunk a" ..
  have "drunk a --> (∀x. drunk x)"
  proof
    assume "drunk a"
    with a show "∀x. drunk x" by (contradiction)
  qed
  then show ?thesis ..
qed

end

lemma deMorgan:

  ¬ (∀x. P x) ==> ∃x. ¬ P x

theorem Drinker's_Principle:

x. drunk x --> (∀x. drunk x)