src/CTT/CTT.thy
changeset 59780 23b67731f4f0
parent 59763 56d2c357e6b5
child 60555 51a6997b1384
--- a/src/CTT/CTT.thy	Mon Mar 23 10:16:20 2015 +0100
+++ b/src/CTT/CTT.thy	Mon Mar 23 13:30:59 2015 +0100
@@ -422,15 +422,15 @@
 
 fun NE_tac ctxt sp i =
   TRY (rtac @{thm EqE} i) THEN
-  Rule_Insts.res_inst_tac ctxt [((("p", 0), Position.none), sp)] @{thm NE} i
+  Rule_Insts.res_inst_tac ctxt [((("p", 0), Position.none), sp)] [] @{thm NE} i
 
 fun SumE_tac ctxt sp i =
   TRY (rtac @{thm EqE} i) THEN
-  Rule_Insts.res_inst_tac ctxt [((("p", 0), Position.none), sp)] @{thm SumE} i
+  Rule_Insts.res_inst_tac ctxt [((("p", 0), Position.none), sp)] [] @{thm SumE} i
 
 fun PlusE_tac ctxt sp i =
   TRY (rtac @{thm EqE} i) THEN
-  Rule_Insts.res_inst_tac ctxt [((("p", 0), Position.none), sp)] @{thm PlusE} i
+  Rule_Insts.res_inst_tac ctxt [((("p", 0), Position.none), sp)] [] @{thm PlusE} i
 
 (** Predicate logic reasoning, WITH THINNING!!  Procedures adapted from NJ. **)