src/HOL/Isar_Examples/Peirce.thy
changeset 63585 f4a308fdf664
parent 61932 2e48182cc82c
equal deleted inserted replaced
63584:68751fe1c036 63585:f4a308fdf664
     3 *)
     3 *)
     4 
     4 
     5 section \<open>Peirce's Law\<close>
     5 section \<open>Peirce's Law\<close>
     6 
     6 
     7 theory Peirce
     7 theory Peirce
     8 imports Main
     8   imports Main
     9 begin
     9 begin
    10 
    10 
    11 text \<open>
    11 text \<open>
    12   We consider Peirce's Law: \<open>((A \<longrightarrow> B) \<longrightarrow> A) \<longrightarrow> A\<close>. This is an inherently
    12   We consider Peirce's Law: \<open>((A \<longrightarrow> B) \<longrightarrow> A) \<longrightarrow> A\<close>. This is an inherently
    13   non-intuitionistic statement, so its proof will certainly involve some form
    13   non-intuitionistic statement, so its proof will certainly involve some form