src/HOL/Tools/Sledgehammer/sledgehammer_atp_translate.ML
Sun, 01 May 2011 18:37:24 +0200 blanchet don't destroy sym table entry for special symbols such as "hAPP" if "explicit_apply" is set
Sun, 01 May 2011 18:37:24 +0200 blanchet cleaned up "explicit_apply" so that it shares most of its code path with the default mode of operation
Sun, 01 May 2011 18:37:24 +0200 blanchet fixed min-arity computation when "explicit_apply" is specified
Sun, 01 May 2011 18:37:24 +0200 blanchet fixed "tags" type encoding after latest round of changes
Sun, 01 May 2011 18:37:24 +0200 blanchet fix handling of proxies after recent drastic changes to the type encodings
Sun, 01 May 2011 18:37:24 +0200 blanchet reconstruct TFF type predicates correctly for ToFoF
Sun, 01 May 2011 18:37:24 +0200 blanchet handle special constants correctly in Isar proof reconstruction code, especially type predicates
Sun, 01 May 2011 18:37:24 +0200 blanchet make sure the minimizer monomorphizes when it should
Sun, 01 May 2011 18:37:24 +0200 blanchet fixed arity of special constants if "explicit_apply" is set
Sun, 01 May 2011 18:37:24 +0200 blanchet make sure typing fact names are unique (needed e.g. by SNARK)
Sun, 01 May 2011 18:37:24 +0200 blanchet minor cleanup
Sun, 01 May 2011 18:37:24 +0200 blanchet reimplemented the hAPP introduction code so that it's done earlier, when the types are still available
Sun, 01 May 2011 18:37:24 +0200 blanchet declare TFF types so that SNARK can be used with types
Sun, 01 May 2011 18:37:24 +0200 blanchet perform constant mangling and/or removal of its type args in an earlier phase, so that the rest of the code doesn't need to worry about it
Sun, 01 May 2011 18:37:24 +0200 blanchet move type declarations to the front, for TFF-compliance
Sun, 01 May 2011 18:37:24 +0200 blanchet use postfix syntax for mangled types, for consistency with unmangled
Sun, 01 May 2011 18:37:24 +0200 blanchet generate typing for "hBOOL" in "Many_Typed" mode + tuning
Sun, 01 May 2011 18:37:24 +0200 blanchet generate pure TFF problems -- ToFoF doesn't like mixtures of FOF and TFF, even when the two logics coincide (e.g. for ground formulas)
Sun, 01 May 2011 18:37:24 +0200 blanchet fake type declarations for full-type args and mangled type encodings, so that type assumptions can be discharged
Sun, 01 May 2011 18:37:24 +0200 blanchet generate TFF type declarations in typed mode
Sun, 01 May 2011 18:37:24 +0200 blanchet added more rudimentary type support to Sledgehammer's ATP encoding
Sun, 01 May 2011 18:37:24 +0200 blanchet fixed type of ATP quantifier types (sic)
Sun, 01 May 2011 18:37:24 +0200 blanchet added "useful_info" argument to ATP formulas -- this will probably be useful later to specify intro, simp, elim to SPASS
Sun, 01 May 2011 18:37:24 +0200 blanchet added support for TFF type declarations
Sun, 01 May 2011 18:37:24 +0200 blanchet reintroduced constructor for formulas, and automatically detect which logic to use (TFF or FOF) to avoid clutter
Sun, 01 May 2011 18:37:24 +0200 blanchet added room for types in ATP quantifiers
Sun, 01 May 2011 18:37:24 +0200 blanchet distinguish FOF and TFF (typed first-order) in ATP abstract syntax tree
Sun, 01 May 2011 18:37:24 +0200 blanchet remove experimental feature ("risky overload")
Sun, 01 May 2011 18:37:24 +0200 blanchet added (without implementation yet) new type encodings for Sledgehammer/ATP
Sun, 01 May 2011 18:37:23 +0200 blanchet close ATP formulas universally earlier, so that we can add type predicates
Sun, 01 May 2011 18:37:23 +0200 blanchet get rid of "explicit_forall" prover-specific option, even if that means some clutter -- foralls will be necessary to attach types to variables
Sun, 01 May 2011 18:37:23 +0200 blanchet renamings
Thu, 21 Apr 2011 22:18:28 +0200 blanchet detect some unsound proofs before showing them to the user
Sat, 16 Apr 2011 16:15:37 +0200 wenzelm modernized structure Proof_Context;
Thu, 14 Apr 2011 11:24:05 +0200 blanchet correctly handle TFrees that occur in (local) facts -- Metis did the right thing here but Sledgehammer was incorrectly generating spurious preconditions such as "dense_linorder(t_a)"
Tue, 05 Apr 2011 11:39:48 +0200 blanchet renamed "const_args" option value to "args"
Tue, 05 Apr 2011 10:54:09 +0200 blanchet temporarily allow useless encoding of helper facts (e.g. fequal_def) instead of throwing exception
Tue, 05 Apr 2011 10:37:21 +0200 blanchet killed unimplemented type encoding "preds"
Tue, 05 Apr 2011 10:37:11 +0200 blanchet remove debugging code
Mon, 04 Apr 2011 18:53:35 +0200 blanchet if "monomorphize" is enabled, mangle the type information in the names by default
Thu, 31 Mar 2011 11:16:52 +0200 blanchet added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
Thu, 17 Mar 2011 11:18:31 +0100 blanchet add option to function to keep trivial ATP formulas, needed for some experiments
Fri, 18 Feb 2011 15:17:39 +0100 blanchet adjust fudge factors
Fri, 18 Feb 2011 12:32:55 +0100 blanchet extended ATP problem syntax to support other applications than Sledgehammer, e.g. experiments with ATPs
Thu, 10 Feb 2011 17:17:31 +0100 blanchet fix handling of "fequal" in generated ATP problems -- the bug was visible if "explicit_apply" was on and "singleton_conv" is one of the facts being translated, as it resulted in an arity error in the ATP (e.g., E)
Tue, 08 Feb 2011 16:10:09 +0100 blanchet sort E weights
Mon, 10 Jan 2011 15:45:46 +0100 wenzelm eliminated Int.toString;
Tue, 28 Dec 2010 18:28:52 +0100 wenzelm made SML/NJ happy;
Wed, 22 Dec 2010 09:02:43 +0100 blanchet made SML/NJ happy
Mon, 20 Dec 2010 12:12:35 +0100 blanchet optionally supply constant weights to E -- turned off by default until properly parameterized
Thu, 16 Dec 2010 15:46:54 +0100 blanchet no need to do a super-duper atomization if Metis fails afterwards anyway
Thu, 16 Dec 2010 15:12:17 +0100 blanchet instantiate induction rules automatically
Wed, 15 Dec 2010 11:26:29 +0100 blanchet tuning
Wed, 15 Dec 2010 11:26:29 +0100 blanchet make sure first-order occurrences of "False" and "True" are handled correctly -- this broke when adding proper support for higher-order occurrences
Wed, 15 Dec 2010 11:26:29 +0100 blanchet consider "finite" overloaded in "precise_overloaded_args" mode
Wed, 15 Dec 2010 11:26:29 +0100 blanchet fix translation of higher-order equality ("fequal") if "precise_overloaded_args" is "true"
Wed, 15 Dec 2010 11:26:28 +0100 blanchet improve partially tagged encoding by adding a helper fact that coalesces consecutive "ti" tags
Wed, 15 Dec 2010 11:26:28 +0100 blanchet added Sledgehammer support for higher-order propositional reasoning
Wed, 15 Dec 2010 11:26:28 +0100 blanchet implemented partially-typed "tags" type encoding
Wed, 15 Dec 2010 11:26:28 +0100 blanchet implemented "no_types" encoding, which is too unsound to be useful but can come in handy for evaluations
less more (0) -60 tip