# HG changeset patch # User blanchet # Date 1335533077 -7200 # Node ID 804fdf0f6006a7a018f01e092d95e896f6a1efc8 # Parent c17cc138064278c1dae37b7d9d34477874fb380b get rid of old CASC setup and move the arithmetic part to a new theory diff -r c17cc1380642 -r 804fdf0f6006 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Fri Apr 27 15:24:37 2012 +0200 +++ b/src/HOL/IsaMakefile Fri Apr 27 15:24:37 2012 +0200 @@ -1142,7 +1142,7 @@ $(LOG)/HOL-TPTP.gz: $(OUT)/HOL \ TPTP/ATP_Problem_Import.thy \ TPTP/ATP_Theory_Export.thy \ - TPTP/CASC_Setup.thy \ + TPTP/THF_Arith.thy \ TPTP/ROOT.ML \ TPTP/TPTP_Parser.thy \ TPTP/TPTP_Parser/ml_yacc_lib.ML \ diff -r c17cc1380642 -r 804fdf0f6006 src/HOL/TPTP/CASC_Setup.thy --- a/src/HOL/TPTP/CASC_Setup.thy Fri Apr 27 15:24:37 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,150 +0,0 @@ -(* Title: HOL/TPTP/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 Complex_Main -uses "sledgehammer_tactics.ML" -begin - -consts - is_int :: "'a \ bool" - is_rat :: "'a \ bool" - -overloading rat_is_int \ "is_int :: rat \ bool" -begin - definition "rat_is_int (q\rat) \ (\n\int. q = of_int n)" -end - -overloading real_is_int \ "is_int :: real \ bool" -begin - definition "real_is_int (x\real) \ x \ \" -end - -overloading real_is_rat \ "is_rat :: real \ bool" -begin - definition "real_is_rat (x\real) \ x \ \" -end - -consts - to_int :: "'a \ int" - to_rat :: "'a \ rat" - to_real :: "'a \ real" - -overloading rat_to_int \ "to_int :: rat \ int" -begin - definition "rat_to_int (q\rat) = floor q" -end - -overloading real_to_int \ "to_int :: real \ int" -begin - definition "real_to_int (x\real) = floor x" -end - -overloading int_to_rat \ "to_rat :: int \ rat" -begin - definition "int_to_rat (n\int) = (of_int n\rat)" -end - -overloading real_to_rat \ "to_rat :: real \ rat" -begin - definition "real_to_rat (x\real) = (inv real x\rat)" -end - -overloading int_to_real \ "to_real :: int \ real" -begin - definition "int_to_real (n\int) = real n" -end - -overloading rat_to_real \ "to_real :: rat \ real" -begin - definition "rat_to_real (x\rat) = (of_rat x\real)" -end - -declare - rat_is_int_def [simp] - real_is_int_def [simp] - real_is_rat_def [simp] - rat_to_int_def [simp] - real_to_int_def [simp] - int_to_rat_def [simp] - real_to_rat_def [simp] - int_to_real_def [simp] - rat_to_real_def [simp] - -lemma to_rat_is_int [intro, simp]: "is_int (to_rat (n\int))" -by (metis int_to_rat_def rat_is_int_def) - -lemma to_real_is_int [intro, simp]: "is_int (to_real (n\int))" -by (metis Ints_real_of_int int_to_real_def real_is_int_def) - -lemma to_real_is_rat [intro, simp]: "is_rat (to_real (q\rat))" -by (metis Rats_of_rat rat_to_real_def real_is_rat_def) - -lemma inj_of_rat [intro, simp]: "inj (of_rat\rat\real)" -by (metis injI of_rat_eq_iff rat_to_real_def) - -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 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 [] - Sledgehammer_Filter.no_relevance_override)) - ORELSE - SOLVE_TIMEOUT (max_secs div 10) "simp" (ALLGOALS (asm_full_simp_tac (simpset_of ctxt))) - ORELSE - SOLVE_TIMEOUT (max_secs div 20) "blast" (ALLGOALS (blast_tac ctxt)) - ORELSE - SOLVE_TIMEOUT (max_secs div 10) "auto" (auto_tac ctxt - THEN ALLGOALS (Sledgehammer_Tactics.sledgehammer_as_oracle_tac ctxt [] - Sledgehammer_Filter.no_relevance_override)) - ORELSE - SOLVE_TIMEOUT (max_secs div 10) "metis" - (ALLGOALS (Metis_Tactic.metis_tac [] ATP_Problem_Generate.liftingN ctxt [])) - ORELSE - SOLVE_TIMEOUT (max_secs div 10) "fast" (ALLGOALS (fast_tac ctxt)) - ORELSE - SOLVE_TIMEOUT (max_secs div 20) "best" (ALLGOALS (best_tac ctxt)) - ORELSE - SOLVE_TIMEOUT (max_secs div 20) "arith" (ALLGOALS (Arith_Data.arith_tac ctxt)) - ORELSE - SOLVE_TIMEOUT (max_secs div 10) "force" (ALLGOALS (force_tac ctxt)) - ORELSE - SOLVE_TIMEOUT max_secs "fastforce" (ALLGOALS (fast_force_tac ctxt)) -*} - -method_setup isabellep = {* - Scan.lift (Scan.optional Parse.nat 6000) >> - (fn m => fn ctxt => SIMPLE_METHOD (isabellep_tac ctxt m)) -*} "combination of Isabelle provers and oracles for CASC" - -end diff -r c17cc1380642 -r 804fdf0f6006 src/HOL/TPTP/ROOT.ML --- a/src/HOL/TPTP/ROOT.ML Fri Apr 27 15:24:37 2012 +0200 +++ b/src/HOL/TPTP/ROOT.ML Fri Apr 27 15:24:37 2012 +0200 @@ -8,11 +8,9 @@ use_thys [ "ATP_Theory_Export", - "TPTP_Interpret" + "TPTP_Interpret", + "THF_Arith" ]; Unsynchronized.setmp Proofterm.proofs (!Proofterm.proofs) - use_thys [ - "ATP_Problem_Import", - "CASC_Setup" - ]; + use_thy "ATP_Problem_Import"; diff -r c17cc1380642 -r 804fdf0f6006 src/HOL/TPTP/THF_Arith.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/TPTP/THF_Arith.thy Fri Apr 27 15:24:37 2012 +0200 @@ -0,0 +1,90 @@ +(* Title: HOL/TPTP/THF_Arith.thy + Author: Jasmin Blanchette + Copyright 2011, 2012 + +Experimental setup for THF arithmetic. This is not connected with the TPTP +parser yet. +*) + +theory THF_Arith +imports Complex_Main +begin + +consts + is_int :: "'a \ bool" + is_rat :: "'a \ bool" + +overloading rat_is_int \ "is_int :: rat \ bool" +begin + definition "rat_is_int (q\rat) \ (\n\int. q = of_int n)" +end + +overloading real_is_int \ "is_int :: real \ bool" +begin + definition "real_is_int (x\real) \ x \ \" +end + +overloading real_is_rat \ "is_rat :: real \ bool" +begin + definition "real_is_rat (x\real) \ x \ \" +end + +consts + to_int :: "'a \ int" + to_rat :: "'a \ rat" + to_real :: "'a \ real" + +overloading rat_to_int \ "to_int :: rat \ int" +begin + definition "rat_to_int (q\rat) = floor q" +end + +overloading real_to_int \ "to_int :: real \ int" +begin + definition "real_to_int (x\real) = floor x" +end + +overloading int_to_rat \ "to_rat :: int \ rat" +begin + definition "int_to_rat (n\int) = (of_int n\rat)" +end + +overloading real_to_rat \ "to_rat :: real \ rat" +begin + definition "real_to_rat (x\real) = (inv real x\rat)" +end + +overloading int_to_real \ "to_real :: int \ real" +begin + definition "int_to_real (n\int) = real n" +end + +overloading rat_to_real \ "to_real :: rat \ real" +begin + definition "rat_to_real (x\rat) = (of_rat x\real)" +end + +declare + rat_is_int_def [simp] + real_is_int_def [simp] + real_is_rat_def [simp] + rat_to_int_def [simp] + real_to_int_def [simp] + int_to_rat_def [simp] + real_to_rat_def [simp] + int_to_real_def [simp] + rat_to_real_def [simp] + +lemma to_rat_is_int [intro, simp]: "is_int (to_rat (n\int))" +by (metis int_to_rat_def rat_is_int_def) + +lemma to_real_is_int [intro, simp]: "is_int (to_real (n\int))" +by (metis Ints_real_of_int int_to_real_def real_is_int_def) + +lemma to_real_is_rat [intro, simp]: "is_rat (to_real (q\rat))" +by (metis Rats_of_rat rat_to_real_def real_is_rat_def) + +lemma inj_of_rat [intro, simp]: "inj (of_rat\rat\real)" +by (metis injI of_rat_eq_iff rat_to_real_def) + +end