blanchet [Fri, 25 Jun 2010 17:26:14 +0200] rev 37580
got rid of "respect_no_atp" option, which even I don't use
blanchet [Fri, 25 Jun 2010 17:13:38 +0200] rev 37579
reorder ML files
blanchet [Fri, 25 Jun 2010 17:08:39 +0200] rev 37578
renamed "Sledgehammer_FOL_Clauses" to "Metis_Clauses", so that Metis doesn't depend on Sledgehammer
blanchet [Fri, 25 Jun 2010 16:42:06 +0200] rev 37577
merge "Sledgehammer_{F,H}OL_Clause", as requested by a FIXME
blanchet [Fri, 25 Jun 2010 16:29:07 +0200] rev 37576
get rid of type alias
blanchet [Fri, 25 Jun 2010 16:27:53 +0200] rev 37575
exploit "Name.desymbolize" to remove some dependencies
blanchet [Fri, 25 Jun 2010 16:15:03 +0200] rev 37574
renamed "Sledgehammer_Fact_Preprocessor" to "Clausifier";
the new name reflects that it's not used only by Sledgehammer (but also by "meson" and "metis") and that it doesn't only clausify facts (but also goals)
blanchet [Fri, 25 Jun 2010 16:03:34 +0200] rev 37573
fewer dependencies
blanchet [Fri, 25 Jun 2010 15:59:13 +0200] rev 37572
more intra-module dependency cleanup + merge "const" and "type_const" tables, since this is safe
blanchet [Fri, 25 Jun 2010 15:30:38 +0200] rev 37571
more moving around of ML files in "Sledgehammer.thy"
blanchet [Fri, 25 Jun 2010 15:22:12 +0200] rev 37570
got rid of needless exception
blanchet [Fri, 25 Jun 2010 15:18:58 +0200] rev 37569
move "MESON" up;
the ultimate goal is to make Sledgehammer depend on MESON and Metis, rather than a big spaghetti
blanchet [Fri, 25 Jun 2010 15:16:22 +0200] rev 37568
remove junk
blanchet [Fri, 25 Jun 2010 15:08:03 +0200] rev 37567
further reduce dependencies on "sledgehammer_fact_filter.ML"
blanchet [Fri, 25 Jun 2010 15:01:35 +0200] rev 37566
move "prepare_clauses" from "sledgehammer_fact_filter.ML" to "sledgehammer_hol_clause.ML";
since it has nothing to do with filtering
wenzelm [Mon, 28 Jun 2010 10:39:39 +0200] rev 37565
merged
Cezary Kaliszyk <kaliszyk@in.tum.de> [Mon, 28 Jun 2010 09:48:36 +0200] rev 37564
Quotient package reverse lifting
Cezary Kaliszyk <kaliszyk@in.tum.de> [Mon, 28 Jun 2010 07:38:39 +0200] rev 37563
Add reverse lifting flag to automated theorem derivation
Cezary Kaliszyk <kaliszyk@in.tum.de> [Mon, 28 Jun 2010 07:32:51 +0200] rev 37562
Restrict quotient definitions to constants
Christian Urban <urbanc@in.tum.de> [Sun, 27 Jun 2010 08:33:01 +0100] rev 37561
mixfix can be given for automatically lifted constants
Christian Urban <urbanc@in.tum.de> [Sat, 26 Jun 2010 08:23:40 +0100] rev 37560
streamlined the generation of quotient theorems out of raw theorems
haftmann [Fri, 25 Jun 2010 19:12:04 +0200] rev 37559
merged
haftmann [Fri, 25 Jun 2010 11:42:29 +0200] rev 37558
avoid REPEAT after THEN_ALL_NEW
wenzelm [Sat, 26 Jun 2010 22:44:25 +0200] rev 37557
refresh Isabelle token marker after buffer properties changed, e.g. when fold mode is switched;
wenzelm [Sat, 26 Jun 2010 22:19:55 +0200] rev 37556
treat alternative newline symbols as in Isabelle/ML;
wenzelm [Sat, 26 Jun 2010 21:26:35 +0200] rev 37555
simplified text_area_painter, with more precise treatment of visible line end;
wenzelm [Fri, 25 Jun 2010 14:05:49 +0200] rev 37554
merged
blanchet [Fri, 25 Jun 2010 12:15:49 +0200] rev 37553
merged
blanchet [Fri, 25 Jun 2010 12:15:24 +0200] rev 37552
eta-expand
blanchet [Fri, 25 Jun 2010 12:08:48 +0200] rev 37551
improve the natural formula relevance filter code, so that it behaves more like the CNF one
blanchet [Fri, 25 Jun 2010 12:07:52 +0200] rev 37550
split SPASS time slot between SOS and non-SOS, in case SOS times out
blanchet [Thu, 24 Jun 2010 21:01:13 +0200] rev 37549
yields ill-typed ATP/metis proofs -- raus!
blanchet [Thu, 24 Jun 2010 21:00:37 +0200] rev 37548
make sure "metisFT" is tried upon "metis" failure in "resolve_inc_tyvars"
haftmann [Fri, 25 Jun 2010 07:19:21 +0200] rev 37547
merged
haftmann [Thu, 24 Jun 2010 21:04:35 +0200] rev 37546
more direct definition simplifies proofs
haftmann [Thu, 24 Jun 2010 21:03:32 +0200] rev 37545
merged
haftmann [Thu, 24 Jun 2010 18:45:31 +0200] rev 37544
more precise tactic: do not escape to a different goal branch (REPEAT is still problematic, though)
blanchet [Thu, 24 Jun 2010 18:22:15 +0200] rev 37543
a76ace919f1c wasn't quite right; second try
blanchet [Thu, 24 Jun 2010 18:04:31 +0200] rev 37542
merge
blanchet [Thu, 24 Jun 2010 17:57:36 +0200] rev 37541
never include anything from the Sledgehammer theory in the relevant facts + killed two obsolete facts
blanchet [Thu, 24 Jun 2010 17:27:18 +0200] rev 37540
better eta-expansion in Sledgehammer's clausification;
make the eta-expansion code more robust in the face of polymorphic arguments + make eta-expansion happen more often, since it first-orderizes things
blanchet [Thu, 24 Jun 2010 17:25:47 +0200] rev 37539
cosmetics
blanchet [Thu, 24 Jun 2010 10:38:01 +0200] rev 37538
make sure that theorems passed using "add:" to Sledgehammer are not eliminated on heuristic grounds
blanchet [Wed, 23 Jun 2010 18:43:50 +0200] rev 37537
improve the new "natural formula" fact filter
wenzelm [Fri, 25 Jun 2010 11:48:37 +0200] rev 37536
explicit treatment of UTF8 sequences as Isabelle symbols;
wenzelm [Thu, 24 Jun 2010 23:20:47 +0200] rev 37535
ML pretty printing of type string according to (quasi-abstract) YXML markup and symbols (including UTF8);
wenzelm [Thu, 24 Jun 2010 22:58:45 +0200] rev 37534
escape UTF8 symbols for the ML compiler;
wenzelm [Thu, 24 Jun 2010 21:57:18 +0200] rev 37533
explicit treatment of UTF8 character sequences as Isabelle symbols;
Christian Urban <urbanc@in.tum.de> [Thu, 24 Jun 2010 16:27:40 +0100] rev 37532
slight cleaning and simplification of the automatic wrapper for quotient definitions
wenzelm [Thu, 24 Jun 2010 17:01:52 +0200] rev 37531
merged
Christian Urban <urbanc@in.tum.de> [Thu, 24 Jun 2010 12:33:51 +0100] rev 37530
export of proper information in the ML-interface of the quotient package
wenzelm [Thu, 24 Jun 2010 14:31:46 +0200] rev 37529
treat Pretty.T as strictly abstract type;
wenzelm [Thu, 24 Jun 2010 14:31:01 +0200] rev 37528
slightly more standard data merge: Symtax.merge (K true) avoids equality on abstract type Pretty.T and gracefully accepts overriding, Symtab.join prefers first entry as usual;
wenzelm [Thu, 24 Jun 2010 14:19:08 +0200] rev 37527
avoid equality on abstract type Pretty.T;
wenzelm [Thu, 24 Jun 2010 13:31:26 +0200] rev 37526
notes on packaging;
wenzelm [Thu, 24 Jun 2010 12:24:35 +0200] rev 37525
misc tuning;
wenzelm [Thu, 24 Jun 2010 12:16:39 +0200] rev 37524
tuned auxiliary structures;
wenzelm [Thu, 24 Jun 2010 11:28:34 +0200] rev 37523
Net.encode_type;
wenzelm [Thu, 24 Jun 2010 11:08:21 +0200] rev 37522
more accurate dependencies;
haftmann [Thu, 24 Jun 2010 09:04:50 +0200] rev 37521
made smlnj happy
blanchet [Wed, 23 Jun 2010 16:28:12 +0200] rev 37520
fix syntax bug in the TPTP output, by ensuring that "hBOOL" is correctly used for n-ary predicates even if (n + k)-ary occurrences of the same predicate, but with a different type, occur in the same problem
blanchet [Wed, 23 Jun 2010 15:35:18 +0200] rev 37519
renamed for easier grep
blanchet [Wed, 23 Jun 2010 15:32:11 +0200] rev 37518
use Skolem inlining also in the nonpolymorphic case, as a step toward simplifying the source code
blanchet [Wed, 23 Jun 2010 15:06:40 +0200] rev 37517
steal some of http://isabelle.in.tum.de/sledgehammer.html and add it to the docs
blanchet [Wed, 23 Jun 2010 14:36:23 +0200] rev 37516
have "metis" method and "metis_tac" fall back on "metisFT" upon failure, following a suggestion by Larry
blanchet [Wed, 23 Jun 2010 12:43:09 +0200] rev 37515
fix bug with "skolem_id" + sort facts for increased readability
blanchet [Wed, 23 Jun 2010 11:36:03 +0200] rev 37514
if SPASS fails at finding a proof with the SOS option turned on, turn it off and try again
blanchet [Wed, 23 Jun 2010 10:20:54 +0200] rev 37513
merged
blanchet [Wed, 23 Jun 2010 10:20:33 +0200] rev 37512
this looks like the most appropriate place to do the beta-eta-contraction
blanchet [Wed, 23 Jun 2010 09:40:06 +0200] rev 37511
killed legacy "neg_clausify" and "clausify"
blanchet [Tue, 22 Jun 2010 23:54:16 +0200] rev 37510
merged
blanchet [Tue, 22 Jun 2010 23:54:02 +0200] rev 37509
factor out TPTP format output into file of its own, to facilitate further changes
blanchet [Tue, 22 Jun 2010 19:10:12 +0200] rev 37508
merged
blanchet [Tue, 22 Jun 2010 19:08:25 +0200] rev 37507
turn on "natural form" filtering in the Mirabelle tests, to see how it performs
blanchet [Tue, 22 Jun 2010 18:47:45 +0200] rev 37506
missing "Unsynchronized" + make exception take a unit
blanchet [Tue, 22 Jun 2010 18:31:49 +0200] rev 37505
added code to optionally perform fact filtering on the original (non-CNF) formulas
blanchet [Tue, 22 Jun 2010 17:10:01 +0200] rev 37504
more cosmetics
blanchet [Tue, 22 Jun 2010 17:07:39 +0200] rev 37503
cosmetics + prevent consideration of inlined Skolem terms in relevance filter
blanchet [Tue, 22 Jun 2010 16:50:55 +0200] rev 37502
canonical argument order
blanchet [Tue, 22 Jun 2010 16:40:36 +0200] rev 37501
leverage new data structure for handling "add:" and "del:"
blanchet [Tue, 22 Jun 2010 16:23:29 +0200] rev 37500
thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet [Tue, 22 Jun 2010 14:48:46 +0200] rev 37499
merge "generic_prover" and "generic_tptp_prover"
blanchet [Tue, 22 Jun 2010 14:28:22 +0200] rev 37498
removed Sledgehammer's support for the DFG syntax;
this removes 350 buggy lines from Sledgehammer. SPASS 3.5 and above support the TPTP syntax.
blanchet [Tue, 22 Jun 2010 13:17:59 +0200] rev 37497
distinguish between "unknown" and "no Kodkodi installed" errors
blanchet [Tue, 22 Jun 2010 13:17:17 +0200] rev 37496
reintroduce new Sledgehammer polymorphic handling code;
this time I tested the proper version of Isabelle
blanchet [Tue, 22 Jun 2010 12:19:06 +0200] rev 37495
make the Nitpick_Example theory processable even when Kodkodi is not installed;
so that at least the "theory" aspects of it (as opposed to the diagnosis offered by Nitpick) are checked on everybody's machines
hoelzl [Wed, 23 Jun 2010 10:05:13 +0200] rev 37494
Make latex happy
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 23 Jun 2010 08:44:44 +0200] rev 37493
Quotient package now uses Partial Equivalence instead place of equivalence
Cezary Kaliszyk <kaliszyk@in.tum.de> [Wed, 23 Jun 2010 08:42:41 +0200] rev 37492
Replace 'list_rel' by 'list_all2'; they are equivalent.
ballarin [Tue, 22 Jun 2010 19:46:16 +0200] rev 37491
Proper treatment of non-inherited mixins.
hoelzl [Tue, 22 Jun 2010 18:15:44 +0200] rev 37490
merged
hoelzl [Mon, 21 Jun 2010 19:33:51 +0200] rev 37489
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
blanchet [Tue, 22 Jun 2010 01:21:52 +0200] rev 37488
reredisable new polymorphic code
blanchet [Mon, 21 Jun 2010 18:45:10 +0200] rev 37487
merged
blanchet [Mon, 21 Jun 2010 18:32:16 +0200] rev 37486
beta-eta was too much, because it transformed SOME x. P x into Eps P, which caused problems later;
reintroduced old proof based on Metis, since it was a good test for the Skolemizer
wenzelm [Mon, 21 Jun 2010 18:31:52 +0200] rev 37485
back to post-release mode;
wenzelm [Mon, 21 Jun 2010 17:41:57 +0200] rev 37484
merged, resolving conflicts in doc-src/IsarRef/Thy/HOL_Specific.thy;
blanchet [Mon, 21 Jun 2010 16:59:37 +0200] rev 37483
make Nitpick's unsound mode a bit more unsound
blanchet [Mon, 21 Jun 2010 14:07:00 +0200] rev 37482
sort cases on the proper key
blanchet [Mon, 21 Jun 2010 13:35:10 +0200] rev 37481
compile
blanchet [Mon, 21 Jun 2010 12:33:43 +0200] rev 37480
thread "full_types"
blanchet [Mon, 21 Jun 2010 12:31:41 +0200] rev 37479
try to improve Sledgehammer/Metis's behavior in full_types mode, e.g. by handing True, False, and If better
blanchet [Mon, 21 Jun 2010 12:28:46 +0200] rev 37478
clean up after fcc768dc9dd0
blanchet [Mon, 21 Jun 2010 11:16:00 +0200] rev 37477
adjusted Nitpick examples to latest changes + make them slightly faster
blanchet [Mon, 21 Jun 2010 11:15:21 +0200] rev 37476
optimized code generated for datatype cases + more;
more = lazy creation of debugging messages in mono code
+ use of "let" when performing some beta-applications (to avoid exponential blowup)
+ removal of some set constructs, to simplify the code and increase precision in some cases (and decrease it in others, but this can be regained)
haftmann [Mon, 21 Jun 2010 09:38:20 +0200] rev 37475
activated Scala code generator test
haftmann [Mon, 21 Jun 2010 09:38:20 +0200] rev 37474
added bot instances; tuned
haftmann [Mon, 21 Jun 2010 09:06:14 +0200] rev 37473
extensionality rule fset_eqI
haftmann [Sun, 20 Jun 2010 22:01:22 +0200] rev 37472
merged
haftmann [Sun, 20 Jun 2010 19:02:41 +0200] rev 37471
separate section for diagnostic commands
haftmann [Sat, 19 Jun 2010 09:50:30 +0200] rev 37470
more binding; avoid arcane Rep and Abs prefixes
haftmann [Sat, 19 Jun 2010 09:14:06 +0200] rev 37469
cleanup of typecopy package
haftmann [Sat, 19 Jun 2010 06:43:33 +0200] rev 37468
quickcheck for fsets
nipkow [Fri, 18 Jun 2010 22:41:16 +0200] rev 37467
merged
nipkow [Fri, 18 Jun 2010 22:40:58 +0200] rev 37466
added pigeonhole lemmas
haftmann [Fri, 18 Jun 2010 21:22:05 +0200] rev 37465
merged
haftmann [Fri, 18 Jun 2010 15:59:51 +0200] rev 37464
tuned whitespace; dropped dead code
haftmann [Fri, 18 Jun 2010 15:26:04 +0200] rev 37463
code_simp: only succeed on real progress
haftmann [Fri, 18 Jun 2010 15:26:02 +0200] rev 37462
prefer fold over foldl
haftmann [Fri, 18 Jun 2010 15:03:21 +0200] rev 37461
conclude simplification with default simpset
haftmann [Fri, 18 Jun 2010 15:03:21 +0200] rev 37460
drop subsumed default equations (requires a little bit unfortunate laziness)
haftmann [Fri, 18 Jun 2010 15:03:20 +0200] rev 37459
avoid Scala legacy operations
haftmann [Fri, 18 Jun 2010 15:03:20 +0200] rev 37458
prefer fold over foldl
haftmann [Fri, 18 Jun 2010 09:21:41 +0200] rev 37457
made List.thy a join point in the theory graph
nipkow [Fri, 18 Jun 2010 20:22:06 +0200] rev 37456
tuned set_replicate lemmas
nipkow [Fri, 18 Jun 2010 14:14:42 +0200] rev 37455
merged
nipkow [Fri, 18 Jun 2010 14:14:29 +0200] rev 37454
added lemmas
haftmann [Fri, 18 Jun 2010 09:04:00 +0200] rev 37453
dropped dead code
haftmann [Thu, 17 Jun 2010 19:32:05 +0200] rev 37452
replaced unreliable metis proof
haftmann [Thu, 17 Jun 2010 16:15:15 +0200] rev 37451
rev is reverse in Haskell
haftmann [Thu, 17 Jun 2010 15:59:48 +0200] rev 37450
first serious draft of a scala code generator
haftmann [Thu, 17 Jun 2010 15:59:47 +0200] rev 37449
more precise code
haftmann [Thu, 17 Jun 2010 15:59:46 +0200] rev 37448
explicit type variable arguments for constructors
haftmann [Thu, 17 Jun 2010 11:33:04 +0200] rev 37447
transitive superclasses were also only a misunderstanding
haftmann [Thu, 17 Jun 2010 10:57:00 +0200] rev 37446
formal introduction of transitive superclasses
haftmann [Thu, 17 Jun 2010 10:51:38 +0200] rev 37445
dropped obscure type argument weakening mapping -- was only a misunderstanding
haftmann [Thu, 17 Jun 2010 10:45:10 +0200] rev 37444
added simp evaluator
haftmann [Thu, 17 Jun 2010 10:02:29 +0200] rev 37443
merged
haftmann [Tue, 15 Jun 2010 14:28:22 +0200] rev 37442
added code_simp infrastructure
haftmann [Tue, 15 Jun 2010 14:28:08 +0200] rev 37441
tuned whitespace
haftmann [Tue, 15 Jun 2010 11:38:40 +0200] rev 37440
maintain cong rules for case combinators; more precise permissiveness
haftmann [Tue, 15 Jun 2010 11:38:40 +0200] rev 37439
drop function definitions of combinators
haftmann [Tue, 15 Jun 2010 11:38:39 +0200] rev 37438
maintain cong rules for case combinators
haftmann [Tue, 15 Jun 2010 08:32:32 +0200] rev 37437
formal introduction of case cong
blanchet [Tue, 15 Jun 2010 16:42:09 +0200] rev 37436
found missing beta-eta-contraction
blanchet [Tue, 15 Jun 2010 16:20:23 +0200] rev 37435
added missing Umlaut
blanchet [Tue, 15 Jun 2010 10:47:06 +0200] rev 37434
make example run a bit faster (might help atbroy102)
haftmann [Tue, 15 Jun 2010 07:42:48 +0200] rev 37433
merged
haftmann [Tue, 15 Jun 2010 07:41:37 +0200] rev 37432
tuned documents
haftmann [Mon, 14 Jun 2010 16:00:47 +0200] rev 37431
teaked naming of superclass projections
haftmann [Mon, 14 Jun 2010 16:00:46 +0200] rev 37430
added lemma funpow_mult
haftmann [Mon, 14 Jun 2010 15:27:11 +0200] rev 37429
extended bib
haftmann [Mon, 14 Jun 2010 15:27:09 +0200] rev 37428
updated generated code
haftmann [Mon, 14 Jun 2010 15:27:09 +0200] rev 37427
added reference
haftmann [Mon, 14 Jun 2010 15:27:08 +0200] rev 37426
subsection on locale interpretation
haftmann [Mon, 14 Jun 2010 12:01:30 +0200] rev 37425
explicitly name and note equations for class eq
haftmann [Mon, 14 Jun 2010 12:01:30 +0200] rev 37424
use various predefined Haskell operations when generating code
haftmann [Mon, 14 Jun 2010 12:01:30 +0200] rev 37423
NEWS
haftmann [Mon, 14 Jun 2010 10:50:49 +0200] rev 37422
tuned internal order
haftmann [Mon, 14 Jun 2010 10:38:29 +0200] rev 37421
dropped unused bindings
haftmann [Mon, 14 Jun 2010 10:38:28 +0200] rev 37420
corrected syntax diagram
blanchet [Mon, 14 Jun 2010 21:49:25 +0200] rev 37419
turn off new polymorphism code again -- a new issue popped up
blanchet [Mon, 14 Jun 2010 20:48:36 +0200] rev 37418
missing case
blanchet [Mon, 14 Jun 2010 20:16:36 +0200] rev 37417
A function called "untyped_aconv" shouldn't look at the bound names!
blanchet [Mon, 14 Jun 2010 19:20:32 +0200] rev 37416
no point in introducing combinators for inlined Skolem functions
blanchet [Mon, 14 Jun 2010 17:12:41 +0200] rev 37415
better error reporting for Vampire
blanchet [Mon, 14 Jun 2010 16:43:44 +0200] rev 37414
expect SPASS 3.7, and give a friendly warning if an older version is used
blanchet [Mon, 14 Jun 2010 16:17:20 +0200] rev 37413
improve ATP-specific error messages
haftmann [Mon, 14 Jun 2010 15:10:50 +0200] rev 37412
merged
haftmann [Mon, 14 Jun 2010 15:10:36 +0200] rev 37411
removed simplifier congruence rule of "prod_case"
blanchet [Mon, 14 Jun 2010 10:36:01 +0200] rev 37410
adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
haftmann [Sat, 12 Jun 2010 15:48:17 +0200] rev 37409
merged
haftmann [Sat, 12 Jun 2010 15:47:50 +0200] rev 37408
declare lexn.simps [code del]
haftmann [Fri, 11 Jun 2010 17:14:02 +0200] rev 37407
declare lex_prod_def [code del]
haftmann [Fri, 11 Jun 2010 17:14:01 +0200] rev 37406
modernized specifications
haftmann [Fri, 11 Jun 2010 17:14:01 +0200] rev 37405
avoid references to old constdefs
blanchet [Sat, 12 Jun 2010 11:12:54 +0200] rev 37404
merged
blanchet [Sat, 12 Jun 2010 11:12:31 +0200] rev 37403
disable new polymorphic code for now, until remaining issues in "equality_inf" are resolved
blanchet [Sat, 12 Jun 2010 11:11:07 +0200] rev 37402
"raise Fail" for internal errors + one new internal error (instead of "Match")
blanchet [Fri, 11 Jun 2010 18:05:05 +0200] rev 37401
make test work again (broken since 09467cdfa198?)
blanchet [Fri, 11 Jun 2010 17:57:16 +0200] rev 37400
adjust Nitpick example to follow latest wave of renamings
blanchet [Fri, 11 Jun 2010 17:10:23 +0200] rev 37399
proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet [Fri, 11 Jun 2010 17:07:27 +0200] rev 37398
beta-eta-contract, to respect "first_order_match"'s specification;
Sledgehammer's Skolem cache sometimes failed without the contraction
blanchet [Fri, 11 Jun 2010 17:05:11 +0200] rev 37397
adjust Nitpick's handling of "<" on "rat"s and "reals"
blanchet [Fri, 11 Jun 2010 16:34:56 +0200] rev 37396
remove needless variables
haftmann [Fri, 11 Jun 2010 16:52:17 +0200] rev 37395
hide sum explicitly
haftmann [Thu, 10 Jun 2010 12:28:27 +0200] rev 37394
merged
haftmann [Thu, 10 Jun 2010 12:26:07 +0200] rev 37393
adjust popular symbolic type constructors
haftmann [Thu, 10 Jun 2010 12:25:14 +0200] rev 37392
tailored set of code equations manually
haftmann [Thu, 10 Jun 2010 12:24:03 +0200] rev 37391
tuned quotes, antiquotations and whitespace
haftmann [Thu, 10 Jun 2010 12:24:02 +0200] rev 37390
moved inductive_codegen to place where product type is available; tuned structure name
haftmann [Thu, 10 Jun 2010 12:24:01 +0200] rev 37389
qualified type "*"; qualified constants Pair, fst, snd, split
haftmann [Tue, 08 Jun 2010 16:37:22 +0200] rev 37388
tuned quotes, antiquotations and whitespace
haftmann [Tue, 08 Jun 2010 16:37:19 +0200] rev 37387
qualified types "+" and nat; qualified constants Ball, Bex, Suc, curry; modernized some specifications
krauss [Thu, 10 Jun 2010 12:08:33 +0200] rev 37386
Adapted Mirabelle script (cf. f60e4dd6d76f)
haftmann [Tue, 08 Jun 2010 07:45:39 +0200] rev 37385
merged
haftmann [Mon, 07 Jun 2010 13:42:38 +0200] rev 37384
more consistent naming aroud type classes and instances
wenzelm [Mon, 07 Jun 2010 17:39:32 +0200] rev 37383
back to non-release mode;
wenzelm [Mon, 21 Jun 2010 11:37:25 +0200] rev 37382
removed obsolete test tags;
wenzelm [Mon, 21 Jun 2010 11:35:56 +0200] rev 37381
Added tag Isabelle2009-2 for changeset 35815ce9218a
wenzelm [Mon, 21 Jun 2010 11:24:19 +0200] rev 37380
final tuning;
haftmann [Mon, 14 Jun 2010 10:38:28 +0200] rev 37379
corrected syntax diagram
krauss [Thu, 10 Jun 2010 12:08:33 +0200] rev 37378
Adapted Mirabelle script (cf. f60e4dd6d76f)
wenzelm [Mon, 14 Jun 2010 21:12:51 +0200] rev 37377
Added tag isa2009-2-test3 for changeset 0eacedd5f780
wenzelm [Mon, 14 Jun 2010 21:10:15 +0200] rev 37376
merged
wenzelm [Fri, 11 Jun 2010 13:25:28 +0200] rev 37375
NEWS: IsabelleText font;
wenzelm [Sun, 13 Jun 2010 23:04:09 +0200] rev 37374
Pretty.string_of (in Scala): actually observe margin/metric;
wenzelm [Sun, 13 Jun 2010 22:33:18 +0200] rev 37373
tuned Command.toString -- preserving uniqueness allows the Scala toplevel to print Linear_Set[Command] results without crashing;
wenzelm [Fri, 11 Jun 2010 21:58:40 +0200] rev 37372
tuned tooltips;
wenzelm [Wed, 09 Jun 2010 18:24:23 +0200] rev 37371
obsolete;
wenzelm [Wed, 09 Jun 2010 16:23:00 +0200] rev 37370
explicit treatment of empty exception block, which could lead to confusing output (e.g. in the theory loader), or even prevent error output altogether;
wenzelm [Wed, 09 Jun 2010 15:09:00 +0200] rev 37369
contrib/README;
wenzelm [Wed, 09 Jun 2010 14:08:08 +0200] rev 37368
removed outdated/confusing INSTALL file;
wenzelm [Tue, 08 Jun 2010 17:45:39 +0200] rev 37367
clarified font_family vs. font_family_default;
install_fonts: refrain from any magic that does not really work on Mac OS, but introduces strange problems on other platforms;
wenzelm [Tue, 08 Jun 2010 13:51:25 +0200] rev 37366
disable set_styles for now -- there are still some race conditions of PropertiesChanged vs. TextArea painting (NB: without it Isabelle_Token_Marker will crash if sub/superscript is actually used);
wenzelm [Mon, 07 Jun 2010 21:48:24 +0200] rev 37365
Added tag isa2009-2-test2 for changeset dfca6c4cd1e8
wenzelm [Mon, 07 Jun 2010 19:21:00 +0200] rev 37364
more uniform treatment of options and attributes, preferring formal markup over old-style LaTeX macros;
wenzelm [Mon, 07 Jun 2010 18:09:18 +0200] rev 37363
merged;
berghofe [Mon, 07 Jun 2010 17:53:02 +0200] rev 37362
Tuned.
berghofe [Mon, 07 Jun 2010 17:52:30 +0200] rev 37361
Documented changes in induct, cases, and nominal_induct method.
wenzelm [Mon, 07 Jun 2010 17:51:26 +0200] rev 37360
merged
wenzelm [Mon, 07 Jun 2010 17:13:36 +0200] rev 37359
made SML/NJ happy again;
wenzelm [Mon, 07 Jun 2010 16:00:35 +0200] rev 37358
recovered some untested theories;
wenzelm [Mon, 07 Jun 2010 17:50:57 +0200] rev 37357
proper target directory;
wenzelm [Mon, 07 Jun 2010 17:50:40 +0200] rev 37356
refer to isabelle-release branch;
wenzelm [Mon, 07 Jun 2010 13:20:05 +0200] rev 37355
no symlinks;
tuned;
wenzelm [Mon, 07 Jun 2010 11:42:54 +0200] rev 37354
merged
wenzelm [Mon, 07 Jun 2010 11:42:42 +0200] rev 37353
tuned ANNOUNCEMENT;
wenzelm [Mon, 07 Jun 2010 11:42:32 +0200] rev 37352
more NEWS;
wenzelm [Mon, 07 Jun 2010 11:27:08 +0200] rev 37351
more NEWS;
tuned;
blanchet [Mon, 07 Jun 2010 10:37:30 +0200] rev 37350
merged
blanchet [Mon, 07 Jun 2010 10:37:06 +0200] rev 37349
cosmetics
blanchet [Sat, 05 Jun 2010 21:30:40 +0200] rev 37348
renaming
blanchet [Sat, 05 Jun 2010 16:39:23 +0200] rev 37347
show more respect for user-specified facts, even if they could lead to unsound proofs + don't throw away "unsound" theorems in "full_type" mode, since they are then sound
blanchet [Sat, 05 Jun 2010 16:08:35 +0200] rev 37346
fix remote Vampire diagnosis
blanchet [Sat, 05 Jun 2010 15:59:58 +0200] rev 37345
make Sledgehammer's "add:" and "del:" syntax work better in the presence of aliases;
some of these aliases pop up only after Sledgehammer has converted the formula to CNF, so it can be very confusing to the user who said "add: foo del: bar" that "bar" is used in the end.
blanchet [Sat, 05 Jun 2010 15:07:50 +0200] rev 37344
totally bypass Sledgehammer's relevance filter when facts are given using the "(fact1 ... factn)" syntax;
faster and more reliable
wenzelm [Sun, 06 Jun 2010 18:47:29 +0200] rev 37343
single heaps archive;
wenzelm [Sun, 06 Jun 2010 18:34:53 +0200] rev 37342
merged
wenzelm [Sun, 06 Jun 2010 18:04:59 +0200] rev 37341
tuned;