src/Pure/Isar/proof.ML
Mon, 25 Oct 2010 21:06:56 +0200 wenzelm renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
Wed, 22 Sep 2010 18:21:48 +0200 wenzelm renamed setmp_noncritical to Unsynchronized.setmp to emphasize its meaning;
Thu, 09 Sep 2010 17:20:27 +0200 wenzelm more abstract treatment of interrupts in structure Exn -- hardly ever need to mention Interrupt literally;
Sat, 04 Sep 2010 00:31:21 +0200 wenzelm recovered options for goal antiquotations from f45d332a90e3: actually pass context to Proof.pretty_goals (see also 45facd8f358e);
Fri, 03 Sep 2010 21:13:53 +0200 wenzelm pretty_goals: turned some global references and function arguments into configuration options (goals_limit = 10, goals_total = true, show_main_goal = false) depending on the context;
Fri, 03 Sep 2010 11:27:35 +0200 wenzelm Proof.pretty_state: print everything in the foreground context to give users a chance to configure options, e.g. via "using [[show_consts]]";
Wed, 25 Aug 2010 21:31:22 +0200 wenzelm added some proof state markup, notably number of subgoals (e.g. for indentation);
Wed, 11 Aug 2010 17:37:04 +0200 wenzelm removed obsolete Proof.get_thmss_cmd (cf. Attrib.eval_thms);
Fri, 06 Aug 2010 14:35:04 +0200 wenzelm removed obsolete Goal.local_future_enforced: Toplevel.run_command is no longer interactive, so Goal.local_future_enabled is sufficient (cf. 349e9223c685 and e07dacec79e7);
Sat, 24 Jul 2010 12:14:53 +0200 wenzelm moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
Sat, 29 May 2010 19:46:29 +0200 wenzelm explicit markup for forked goals, as indicated by Goal.fork;
Thu, 29 Apr 2010 17:29:53 +0200 wenzelm 'write': actually observe the proof structure (like 'let' or 'fix');
Thu, 29 Apr 2010 16:55:22 +0200 wenzelm ProofContext.read_const: allow for type constraint (for fixed variable);
Mon, 26 Apr 2010 20:30:50 +0200 wenzelm eliminanated some unreferenced identifiers;
Sun, 25 Apr 2010 16:10:05 +0200 wenzelm removed obsolete/unused Proof.match_bind;
Sun, 25 Apr 2010 15:52:03 +0200 wenzelm modernized naming conventions of main Isar proof elements;
Sun, 25 Apr 2010 15:13:33 +0200 wenzelm goals: simplified handling of implicit variables -- removed obsolete warning;
Sun, 07 Mar 2010 11:57:16 +0100 wenzelm modernized structure Local_Defs;
Fri, 15 Jan 2010 13:37:41 +0100 berghofe Eliminated is_open option of Rule_Cases.make_nested/make_common;
Mon, 04 Jan 2010 15:35:53 +0100 wenzelm after_qed: refrain from Position.setmp_thread_data, which causes duplication of results with several independent proof attempts;
Mon, 02 Nov 2009 20:57:48 +0100 wenzelm modernized structure Proof_Display;
Mon, 02 Nov 2009 20:45:23 +0100 wenzelm modernized structure AutoBind;
Sun, 01 Nov 2009 15:24:45 +0100 wenzelm modernized structure Rule_Cases;
Wed, 28 Oct 2009 22:01:05 +0100 wenzelm replaced slightly odd get_goal/flat_goal by explicit goal views that correspond to the usual method categories;
Sun, 25 Oct 2009 19:17:42 +0100 wenzelm more direct access to naming;
Tue, 20 Oct 2009 21:26:45 +0200 wenzelm backpatching of structure Proof and ProofContext -- avoid odd aliases;
Fri, 02 Oct 2009 21:39:06 +0200 wenzelm clarified Proof.refine_insert -- always "refine" to apply standard method treatment (of conjunctions);
Wed, 30 Sep 2009 22:25:50 +0200 wenzelm eliminated dead code;
Wed, 30 Sep 2009 00:57:28 +0200 wenzelm replaced chained_goal by slightly more appropriate flat_goal;
Wed, 30 Sep 2009 00:17:06 +0200 wenzelm added chained_goal, which presents the goal thm as seen by semi-structured methods;
Tue, 29 Sep 2009 11:49:22 +0200 wenzelm explicit indication of Unsynchronized.ref;
Sat, 25 Jul 2009 18:02:43 +0200 wenzelm basic method application: avoid Position.setmp_thread_data_seq, which destroys transaction context;
Sat, 25 Jul 2009 10:31:27 +0200 wenzelm renamed structure Display_Goal to Goal_Display;
Fri, 24 Jul 2009 11:50:35 +0200 wenzelm Display_Goal.pretty_goals: always Markup.subgoal, clarified options;
Thu, 23 Jul 2009 16:52:16 +0200 wenzelm clarified pretty_goals, pretty_thm_aux: plain context;
Tue, 21 Jul 2009 01:03:18 +0200 wenzelm proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
Mon, 20 Jul 2009 21:20:09 +0200 wenzelm moved pretty_goals etc. to Display_Goal (required by tracing tacticals);
Mon, 20 Jul 2009 00:37:39 +0200 wenzelm Proof.future_proof: declare all assumptions as well;
Sun, 19 Jul 2009 19:24:04 +0200 wenzelm parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
Sun, 19 Jul 2009 18:02:40 +0200 wenzelm more abstract Future.is_worker;
Sat, 28 Mar 2009 17:53:33 +0100 wenzelm renamed ProofContext.add_fixes_i to ProofContext.add_fixes, eliminated obsolete external version;
Sat, 28 Mar 2009 17:21:11 +0100 wenzelm renamed ProofContext.note_thmss_i to ProofContext.note_thmss, eliminated obsolete external version;
Sat, 28 Mar 2009 17:10:43 +0100 wenzelm simplified references to facts, eliminated external note_thmss;
Sat, 28 Mar 2009 16:31:16 +0100 wenzelm replaced add_binds(_i) by bind_terms -- internal version only;
Tue, 17 Mar 2009 13:33:21 +0100 wenzelm goal_tac: finish marked assumptions from left to right -- corresponds better with the strategy of etac, with significant performance gains in some situations;
Fri, 13 Mar 2009 19:58:26 +0100 wenzelm unified type Proof.method and pervasive METHOD combinators;
Tue, 10 Mar 2009 21:18:52 +0100 wenzelm invoke_case: proper qualification of name binding, avoiding old no_base_names;
Thu, 05 Mar 2009 11:58:53 +0100 wenzelm eliminated obsolete ProofContext.full_bname;
Tue, 03 Mar 2009 14:07:43 +0100 wenzelm Thm.binding;
Wed, 21 Jan 2009 16:47:32 +0100 haftmann binding is alias for Binding.T
Sun, 11 Jan 2009 18:18:35 +0100 wenzelm added Goal.future_enabled abstraction -- now also checks that this is already
Sat, 10 Jan 2009 21:32:30 +0100 wenzelm added parallel_proofs flag (default true, cf. usedir option -Q), which can be disabled in low-memory situations;
Thu, 08 Jan 2009 13:18:34 +0100 wenzelm tuned;
Thu, 08 Jan 2009 08:06:11 +0100 haftmann made SML/NJ happy
Wed, 07 Jan 2009 20:27:05 +0100 wenzelm future_proof: refined version covers local_future_proof and global_future_proof;
Wed, 07 Jan 2009 16:22:10 +0100 wenzelm qed/after_qed: singleton result;
Wed, 07 Jan 2009 12:09:39 +0100 wenzelm future_terminal_proof: no fork for interactive mode, assert_backward;
Tue, 06 Jan 2009 13:36:42 +0100 wenzelm future_terminal_proof: check Future.enabled;
Mon, 05 Jan 2009 14:22:40 +0100 wenzelm added future_terminal_proof;
Sun, 04 Jan 2009 15:40:30 +0100 wenzelm more precise is_relevant: requires original main goal, not initial goal state;
Fri, 12 Dec 2008 22:13:13 +0100 wenzelm global_qed: refrain from ProofContext.auto_bind_facts, to avoid
Fri, 05 Dec 2008 00:23:37 +0100 wenzelm merged
Thu, 04 Dec 2008 23:00:58 +0100 wenzelm future proofs: pass actual futures to facilitate composite computations;
Thu, 04 Dec 2008 14:43:33 +0100 haftmann cleaned up binding module and related code
Thu, 16 Oct 2008 22:44:36 +0200 wenzelm conclude_goal: precise goal context, include all sorts from context into statement, check shyps of result;
Thu, 02 Oct 2008 14:22:40 +0200 wenzelm simplified Exn.EXCEPTIONS;
Wed, 01 Oct 2008 22:33:24 +0200 wenzelm replaced can_defer by is_relevant (negation);
Wed, 01 Oct 2008 18:16:14 +0200 wenzelm future_proof: protect conclusion of deferred proof state;
Wed, 01 Oct 2008 12:18:18 +0200 wenzelm renamed promise to future, tuned related interfaces;
Wed, 01 Oct 2008 12:00:02 +0200 wenzelm more robust treatment of Interrupt (cf. exn.ML);
Tue, 30 Sep 2008 23:31:38 +0200 wenzelm promise_proof: proper statement with empty vars;
Mon, 29 Sep 2008 21:26:32 +0200 wenzelm promise global proofs;
Fri, 26 Sep 2008 19:07:56 +0200 wenzelm eliminated polymorphic equality;
Thu, 18 Sep 2008 14:06:58 +0200 wenzelm show: non-critical testing;
Wed, 17 Sep 2008 22:06:57 +0200 wenzelm added map_contexts;
Tue, 02 Sep 2008 16:55:33 +0200 wenzelm type Attrib.binding abbreviates Name.binding without attributes;
Tue, 02 Sep 2008 14:10:45 +0200 wenzelm explicit type Name.binding for higher-specification elements;
Fri, 16 May 2008 21:53:27 +0200 wenzelm removed obsolete case rule_context;
Wed, 19 Mar 2008 22:27:57 +0100 wenzelm renamed datatype thmref to Facts.ref, tuned interfaces;
Tue, 11 Mar 2008 17:13:41 +0100 wenzelm put_facts: do_props, i.e. facts are indexed by proposition again;
Thu, 24 Jan 2008 23:51:19 +0100 wenzelm statement: keep explicit position;
Mon, 29 Oct 2007 16:13:47 +0100 wenzelm test_proof: do not change Proofterm.proofs here (not thread-safe);
Thu, 11 Oct 2007 16:05:30 +0200 wenzelm moved ProofContext.pp to Syntax.pp;
Tue, 09 Oct 2007 00:20:13 +0200 wenzelm generic Syntax.pretty/string_of operations;
Mon, 01 Oct 2007 15:14:56 +0200 wenzelm ContextPosition.put_ctxt;
Fri, 07 Sep 2007 22:13:48 +0200 wenzelm reset goal messages after goal update;
Fri, 07 Sep 2007 09:11:32 +0200 wenzelm made smlnj happy;
Thu, 06 Sep 2007 12:30:41 +0200 wenzelm added goal_message;
Sun, 29 Jul 2007 14:30:06 +0200 wenzelm renamed Drule.is_dummy_thm to Thm.is_dummy;
Fri, 27 Jul 2007 21:55:21 +0200 wenzelm chain/using: filter out dummy_thm;
Tue, 24 Jul 2007 19:44:35 +0200 wenzelm moved exception capture/release to structure Exn;
Fri, 13 Jul 2007 22:36:10 +0200 wenzelm tuned;
Wed, 11 Jul 2007 12:01:10 +0200 berghofe Proof terms for meta-conjunctions are now normalized before
Sat, 07 Jul 2007 18:39:19 +0200 wenzelm pretty_state: subgoal markup;
Tue, 19 Jun 2007 23:15:27 +0200 wenzelm balanced conjunctions;
Wed, 13 Jun 2007 00:02:03 +0200 wenzelm apply_method/end_proof: pass position;
Sat, 09 Dec 2006 18:05:47 +0100 wenzelm init_context: reset naming;
Thu, 07 Dec 2006 00:42:04 +0100 wenzelm reorganized structure Goal vs. Tactic;
Tue, 05 Dec 2006 22:14:51 +0100 wenzelm generic_goal: enable stmt mode;
Tue, 28 Nov 2006 00:35:18 +0100 wenzelm simplified '?' operator;
Wed, 22 Nov 2006 15:58:59 +0100 wenzelm init: enter inner statement mode, which prevents local notes from being named internally;
Tue, 21 Nov 2006 20:58:15 +0100 wenzelm made SML/NJ happy;
Tue, 21 Nov 2006 20:48:06 +0100 wenzelm removed obsolete simple_note_thms;
Tue, 21 Nov 2006 18:07:40 +0100 wenzelm simplified theorem(_i);
Tue, 14 Nov 2006 15:29:56 +0100 wenzelm simplified Proof.theorem(_i) interface -- removed target support;
Thu, 09 Nov 2006 21:44:32 +0100 wenzelm Stack.map_top;
Mon, 09 Oct 2006 02:20:06 +0200 wenzelm def(_i): LocalDefs.add_defs;
Sat, 07 Oct 2006 01:30:58 +0200 wenzelm tuned;
Wed, 09 Aug 2006 00:12:33 +0200 wenzelm global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
Sat, 05 Aug 2006 14:52:55 +0200 wenzelm removed obsolete sign_of;
Wed, 02 Aug 2006 22:27:03 +0200 wenzelm simplified Assumption/ProofContext.export;
Thu, 27 Jul 2006 13:43:01 +0200 wenzelm moved basic assumption operations from structure ProofContext to Assumption;
Wed, 26 Jul 2006 00:44:46 +0200 wenzelm export goal_tac (was internal refine_tac);
Thu, 06 Jul 2006 17:47:35 +0200 wenzelm apply_text: support Method.Source_i;
Tue, 04 Jul 2006 18:39:58 +0200 wenzelm added schematic_goal;
Sat, 17 Jun 2006 19:37:57 +0200 wenzelm ProofContext.exports: simultaneous facts;
Thu, 15 Jun 2006 23:10:45 +0200 wenzelm ProofContext: moved variable operations to struct Variable;
Mon, 12 Jun 2006 21:19:02 +0200 wenzelm Unify.matches_list;
Sun, 11 Jun 2006 21:59:26 +0200 wenzelm fixes: include mixfix syntax;
Mon, 05 Jun 2006 21:54:20 +0200 wenzelm allow non-trivial schematic goals (via embedded term vars);
less more (0) -120 tip