src/HOL/Induct/PropLog.thy
changeset 9101 b643f4d7b9e9
parent 5184 9b8547a9496a
child 10759 994877ee68cb
--- a/src/HOL/Induct/PropLog.thy	Wed Jun 21 15:58:23 2000 +0200
+++ b/src/HOL/Induct/PropLog.thy	Wed Jun 21 18:09:09 2000 +0200
@@ -6,7 +6,7 @@
 Inductive definition of propositional logic.
 *)
 
-PropLog = Finite + Datatype +
+PropLog = Main +
 
 datatype
     'a pl = false | var 'a ("#_" [1000]) | "->" ('a pl) ('a pl) (infixr 90)
@@ -15,7 +15,7 @@
   "|-"  :: ['a pl set, 'a pl] => bool   (infixl 50)
   "|="  :: ['a pl set, 'a pl] => bool   (infixl 50)
   eval2 :: ['a pl, 'a set] => bool
-  eval  :: ['a set, 'a pl] => bool      ("_[_]" [100,0] 100)
+  eval  :: ['a set, 'a pl] => bool      ("_[[_]]" [100,0] 100)
   hyps  :: ['a pl, 'a set] => 'a pl set
 
 translations
@@ -30,8 +30,8 @@
   MP  "[| H |- p->q; H |- p |] ==> H |- q"
 
 defs
-  sat_def  "H |= p  ==  (!tt. (!q:H. tt[q]) --> tt[p])"
-  eval_def "tt[p] == eval2 p tt"
+  sat_def  "H |= p  ==  (!tt. (!q:H. tt[[q]]) --> tt[[p]])"
+  eval_def "tt[[p]] == eval2 p tt"
 
 primrec
   "eval2(false) = (%x. False)"