src/ZF/Induct/PropLog.thy
changeset 65449 c82e63b11b8b
parent 61798 27f3c10b0b50
child 69587 53982d5ec0bb
equal deleted inserted replaced
65448:9bc3b57c1fa7 65449:c82e63b11b8b
     3     Copyright   1993  University of Cambridge
     3     Copyright   1993  University of Cambridge
     4 *)
     4 *)
     5 
     5 
     6 section \<open>Meta-theory of propositional logic\<close>
     6 section \<open>Meta-theory of propositional logic\<close>
     7 
     7 
     8 theory PropLog imports Main begin
     8 theory PropLog imports ZF begin
     9 
     9 
    10 text \<open>
    10 text \<open>
    11   Datatype definition of propositional logic formulae and inductive
    11   Datatype definition of propositional logic formulae and inductive
    12   definition of the propositional tautologies.
    12   definition of the propositional tautologies.
    13 
    13