src/HOL/ex/Hilbert_Classical.thy
changeset 16417 9bc16273c2d4
parent 14981 e73f8140af78
child 26828 60d1fa8e0831
equal deleted inserted replaced
16416:6061ae1f90f2 16417:9bc16273c2d4
     3     Author:     Stefan Berghofer and Markus Wenzel, TU Muenchen
     3     Author:     Stefan Berghofer and Markus Wenzel, TU Muenchen
     4 *)
     4 *)
     5 
     5 
     6 header {* Hilbert's choice and classical logic *}
     6 header {* Hilbert's choice and classical logic *}
     7 
     7 
     8 theory Hilbert_Classical = Main:
     8 theory Hilbert_Classical imports Main begin
     9 
     9 
    10 text {*
    10 text {*
    11   Derivation of the classical law of tertium-non-datur by means of
    11   Derivation of the classical law of tertium-non-datur by means of
    12   Hilbert's choice operator (due to M. J. Beeson and J. Harrison).
    12   Hilbert's choice operator (due to M. J. Beeson and J. Harrison).
    13 *}
    13 *}