--- a/src/HOL/IsaMakefile Mon May 02 10:50:09 2011 +0200
+++ b/src/HOL/IsaMakefile Mon May 02 12:09:33 2011 +0200
@@ -1035,8 +1035,8 @@
$(LOG)/HOL-ex.gz: $(OUT)/HOL Decision_Procs/Commutative_Ring.thy \
Number_Theory/Primes.thy ex/Abstract_NAT.thy ex/Antiquote.thy \
ex/Arith_Examples.thy ex/Arithmetic_Series_Complex.thy ex/BT.thy \
- ex/BinEx.thy ex/Binary.thy ex/Birthday_Paradoxon.thy ex/CTL.thy \
- ex/Case_Product.thy \
+ ex/BinEx.thy ex/Binary.thy ex/Birthday_Paradoxon.thy \
+ ex/CASC_Setup.thy ex/CTL.thy ex/Case_Product.thy \
ex/Chinese.thy ex/Classical.thy ex/CodegenSML_Test.thy \
ex/Coercion_Examples.thy ex/Coherent.thy ex/Dedekind_Real.thy \
ex/Efficient_Nat_examples.thy ex/Eval_Examples.thy ex/Fundefs.thy \
@@ -1056,7 +1056,7 @@
ex/SVC_Oracle.thy ex/Serbian.thy ex/Set_Algebras.thy \
ex/sledgehammer_tactics.ML ex/Sqrt.thy \
ex/Sqrt_Script.thy ex/Sudoku.thy ex/Tarski.thy ex/Termination.thy \
- ex/TPTP.thy ex/Transfer_Ex.thy ex/Tree23.thy ex/Unification.thy \
+ ex/Transfer_Ex.thy ex/Tree23.thy ex/Unification.thy \
ex/While_Combinator_Example.thy ex/document/root.bib \
ex/document/root.tex ex/set.thy ex/svc_funcs.ML ex/svc_test.thy \
../Tools/interpretation_with_defs.ML
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/CASC_Setup.thy Mon May 02 12:09:33 2011 +0200
@@ -0,0 +1,65 @@
+(* Title: HOL/ex/CASC_Setup.thy
+ Author: Jasmin Blanchette
+ Copyright 2011
+
+Setup for Isabelle, Nitpick, and Refute for participating at CASC in the THF and
+TNT divisions. This theory file should be loaded by the Isabelle theory files
+generated by Geoff Sutcliffe's TPTP2X tool from the original THF0 files.
+*)
+
+theory CASC_Setup
+imports Main
+uses "sledgehammer_tactics.ML"
+begin
+
+declare mem_def [simp add]
+
+declare [[smt_oracle]]
+
+refute_params [maxtime = 10000, no_assms, expect = genuine]
+nitpick_params [timeout = none, card = 1-50, verbose, dont_box, no_assms,
+ batch_size = 1, expect = genuine]
+
+ML {* Proofterm.proofs := 0 *}
+
+ML {*
+fun SOLVE_TIMEOUT seconds name tac st =
+ let
+ val result =
+ TimeLimit.timeLimit (Time.fromSeconds seconds)
+ (fn () => SINGLE (SOLVE tac) st) ()
+ handle TimeLimit.TimeOut => NONE
+ | ERROR _ => NONE
+ in
+ (case result of
+ NONE => (warning ("FAILURE: " ^ name); Seq.empty)
+ | SOME st' => (warning ("SUCCESS: " ^ name); Seq.single st'))
+ end
+*}
+
+ML {*
+fun isabellep_tac ctxt cs ss css max_secs =
+ SOLVE_TIMEOUT (max_secs div 10) "smt" (ALLGOALS (SMT_Solver.smt_tac ctxt []))
+ ORELSE
+ SOLVE_TIMEOUT (max_secs div 5) "sledgehammer"
+ (ALLGOALS (Sledgehammer_Tactics.sledgehammer_as_oracle_tac ctxt))
+ ORELSE
+ SOLVE_TIMEOUT (max_secs div 10) "simp" (ALLGOALS (asm_full_simp_tac ss))
+ ORELSE
+ SOLVE_TIMEOUT (max_secs div 20) "blast" (ALLGOALS (blast_tac cs))
+ ORELSE
+ SOLVE_TIMEOUT (max_secs div 10) "auto" (auto_tac css
+ THEN ALLGOALS (Sledgehammer_Tactics.sledgehammer_as_oracle_tac ctxt))
+ ORELSE
+ SOLVE_TIMEOUT (max_secs div 10) "metis" (ALLGOALS (Metis_Tactics.metis_tac ctxt []))
+ ORELSE
+ SOLVE_TIMEOUT (max_secs div 10) "fast" (ALLGOALS (fast_tac cs))
+ ORELSE
+ SOLVE_TIMEOUT (max_secs div 10) "best" (ALLGOALS (best_tac cs))
+ ORELSE
+ SOLVE_TIMEOUT (max_secs div 10) "force" (ALLGOALS (force_tac css))
+ ORELSE
+ SOLVE_TIMEOUT max_secs "fastsimp" (ALLGOALS (fast_simp_tac css))
+*}
+
+end
--- a/src/HOL/ex/ROOT.ML Mon May 02 10:50:09 2011 +0200
+++ b/src/HOL/ex/ROOT.ML Mon May 02 12:09:33 2011 +0200
@@ -78,7 +78,7 @@
];
Unsynchronized.setmp Proofterm.proofs (! Proofterm.proofs)
- use_thy "TPTP";
+ use_thy "CASC_THF0";
if getenv "ISABELLE_GHC" = "" then ()
else use_thy "Quickcheck_Narrowing_Examples";
--- a/src/HOL/ex/TPTP.thy Mon May 02 10:50:09 2011 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,63 +0,0 @@
-(* Title: HOL/ex/TPTP.thy
- Author: Jasmin Blanchette
- Copyright 2011
-
-TPTP "IsabelleP" tactic.
-*)
-
-theory TPTP
-imports Main
-uses "sledgehammer_tactics.ML"
-begin
-
-declare mem_def [simp add]
-
-declare [[smt_oracle]]
-
-refute_params [maxtime = 10000, no_assms, expect = genuine]
-nitpick_params [timeout = none, card = 1-50, verbose, dont_box, no_assms,
- batch_size = 1, expect = genuine]
-
-ML {* Proofterm.proofs := 0 *}
-
-ML {*
-fun SOLVE_TIMEOUT seconds name tac st =
- let
- val result =
- TimeLimit.timeLimit (Time.fromSeconds seconds)
- (fn () => SINGLE (SOLVE tac) st) ()
- handle TimeLimit.TimeOut => NONE
- | ERROR _ => NONE
- in
- (case result of
- NONE => (warning ("FAILURE: " ^ name); Seq.empty)
- | SOME st' => (warning ("SUCCESS: " ^ name); Seq.single st'))
- end
-*}
-
-ML {*
-fun isabellep_tac ctxt cs ss css max_secs =
- SOLVE_TIMEOUT (max_secs div 10) "smt" (ALLGOALS (SMT_Solver.smt_tac ctxt []))
- ORELSE
- SOLVE_TIMEOUT (max_secs div 5) "sledgehammer"
- (ALLGOALS (Sledgehammer_Tactics.sledgehammer_as_oracle_tac ctxt))
- ORELSE
- SOLVE_TIMEOUT (max_secs div 10) "simp" (ALLGOALS (asm_full_simp_tac ss))
- ORELSE
- SOLVE_TIMEOUT (max_secs div 20) "blast" (ALLGOALS (blast_tac cs))
- ORELSE
- SOLVE_TIMEOUT (max_secs div 10) "auto" (auto_tac css
- THEN ALLGOALS (Sledgehammer_Tactics.sledgehammer_as_oracle_tac ctxt))
- ORELSE
- SOLVE_TIMEOUT (max_secs div 10) "metis" (ALLGOALS (Metis_Tactics.metis_tac ctxt []))
- ORELSE
- SOLVE_TIMEOUT (max_secs div 10) "fast" (ALLGOALS (fast_tac cs))
- ORELSE
- SOLVE_TIMEOUT (max_secs div 10) "best" (ALLGOALS (best_tac cs))
- ORELSE
- SOLVE_TIMEOUT (max_secs div 10) "force" (ALLGOALS (force_tac css))
- ORELSE
- SOLVE_TIMEOUT max_secs "fastsimp" (ALLGOALS (fast_simp_tac css))
-*}
-
-end