src/HOL/Induct/PropLog.thy
changeset 16417 9bc16273c2d4
parent 13075 d3e1d554cd6d
child 16563 a92f96951355
equal deleted inserted replaced
16416:6061ae1f90f2 16417:9bc16273c2d4
     4     Copyright   1994  TU Muenchen & University of Cambridge
     4     Copyright   1994  TU Muenchen & University of Cambridge
     5 *)
     5 *)
     6 
     6 
     7 header {* Meta-theory of propositional logic *}
     7 header {* Meta-theory of propositional logic *}
     8 
     8 
     9 theory PropLog = Main:
     9 theory PropLog imports Main begin
    10 
    10 
    11 text {*
    11 text {*
    12   Datatype definition of propositional logic formulae and inductive
    12   Datatype definition of propositional logic formulae and inductive
    13   definition of the propositional tautologies.
    13   definition of the propositional tautologies.
    14 
    14