src/Pure/Isar/toplevel.ML
Thu, 24 Jan 2008 23:51:22 +0100 wenzelm removed unused properties;
Sun, 06 Jan 2008 15:57:54 +0100 wenzelm removed obsolete prompt markup;
Thu, 03 Jan 2008 22:25:15 +0100 wenzelm replaced thread_properties by simplified version in position.ML;
Thu, 03 Jan 2008 17:50:43 +0100 wenzelm toplevel print_exn: proper setmp_thread_properties;
Thu, 03 Jan 2008 00:15:42 +0100 wenzelm maintain thread transition properties;
Wed, 02 Jan 2008 23:00:57 +0100 wenzelm type transition: added properties field;
Sat, 08 Dec 2007 21:20:07 +0100 wenzelm Isar loop: recover after toplevel crashes;
Thu, 06 Dec 2007 00:21:30 +0100 wenzelm replaced Markup.enclose by Markup.markup, which operates on plain strings instead of raw output;
Tue, 04 Dec 2007 22:49:28 +0100 wenzelm Toplevel.loop: explicit argument for secure loop, no warning on quit;
Sun, 18 Nov 2007 16:10:11 +0100 wenzelm init_empty: check before change (avoids non-linear update);
Mon, 05 Nov 2007 20:50:45 +0100 wenzelm simplified LocalTheory.reinit;
Fri, 02 Nov 2007 18:53:00 +0100 haftmann clarified theory target interface
Sun, 28 Oct 2007 11:57:04 +0100 wenzelm safe_exit: controlled_execution;
Tue, 09 Oct 2007 17:11:20 +0200 wenzelm TheoryTarget.init_cmd;
Tue, 09 Oct 2007 00:20:13 +0200 wenzelm generic Syntax.pretty/string_of operations;
Mon, 01 Oct 2007 15:14:57 +0200 wenzelm print_state_context: local theory context, not proof context;
Sun, 30 Sep 2007 16:20:42 +0200 wenzelm local_theory transactions: more careful treatment of context position;
Tue, 18 Sep 2007 18:05:37 +0200 wenzelm simplified PrintMode interfaces;
Tue, 28 Aug 2007 18:04:21 +0200 berghofe Added local_theory_to_proof'
Thu, 16 Aug 2007 18:53:21 +0200 wenzelm global state transformation: non-critical, but also non-thread-safe;
Sun, 29 Jul 2007 17:28:55 +0200 wenzelm NAMED_CRITICAL;
Sun, 29 Jul 2007 16:00:04 +0200 wenzelm added TOPLEVEL_ERROR (simplified version from output.ML);
Tue, 24 Jul 2007 22:59:53 +0200 wenzelm *** empty log message ***
Tue, 24 Jul 2007 19:44:35 +0200 wenzelm moved exception capture/release to structure Exn;
Mon, 23 Jul 2007 19:45:47 +0200 wenzelm marked some CRITICAL sections;
Sun, 22 Jul 2007 21:20:56 +0200 wenzelm init_empty: invoke operation *after* safe_exit;
Tue, 10 Jul 2007 23:29:44 +0200 wenzelm Markup.enclose;
Tue, 10 Jul 2007 16:45:04 +0200 wenzelm Markup.output;
Mon, 09 Jul 2007 11:44:22 +0200 wenzelm prompt: plain string, not output;
Sun, 08 Jul 2007 13:10:57 +0200 wenzelm simplified/more robust print_state;
Sat, 07 Jul 2007 18:39:20 +0200 wenzelm toplevel prompt/print_state: proper markup, removed hooks;
Sat, 07 Jul 2007 00:15:00 +0200 wenzelm tuned;
Tue, 19 Jun 2007 23:21:39 +0200 wenzelm oops -- fixed profiling;
Tue, 19 Jun 2007 23:15:59 +0200 wenzelm profiling: observe no_timing;
Wed, 13 Jun 2007 00:02:06 +0200 wenzelm transaction: context_position (non-destructive only);
Wed, 04 Apr 2007 00:11:21 +0200 wenzelm removed unused info channel;
Sat, 20 Jan 2007 14:09:20 +0100 wenzelm Toplevel.debug: coincide with Output.debugging;
Fri, 19 Jan 2007 22:08:02 +0100 wenzelm moved ML context stuff to from Context to ML_Context;
Fri, 19 Jan 2007 13:09:37 +0100 wenzelm added generic_theory_of;
Wed, 10 Jan 2007 20:16:52 +0100 wenzelm fixed exit: proper type check of state;
Fri, 05 Jan 2007 14:30:08 +0100 wenzelm removed Toplevel.print_exn hook -- existing error_msg_fn does the job;
Thu, 04 Jan 2007 11:56:53 +0100 haftmann eliminated Option.app
Wed, 03 Jan 2007 21:09:13 +0100 aspinall Expose command failure recovery code for top level.
Sat, 30 Dec 2006 16:08:00 +0100 wenzelm removed conditional combinator;
Sat, 30 Dec 2006 12:33:28 +0100 wenzelm refined notion of empty toplevel, admits undo of 'end';
Fri, 15 Dec 2006 00:08:16 +0100 wenzelm removed obsolete assert;
Thu, 23 Nov 2006 22:38:30 +0100 wenzelm prefer Proof.context over Context.generic;
Sat, 11 Nov 2006 21:13:12 +0100 wenzelm level: do not account for local theory blocks (relevant for document preparation);
Fri, 10 Nov 2006 22:18:54 +0100 wenzelm simplified local theory wrappers;
Thu, 09 Nov 2006 21:44:35 +0100 wenzelm LocalTheory.restore;
Tue, 07 Nov 2006 19:39:53 +0100 wenzelm removed obsolete print_state_hook;
Sat, 04 Nov 2006 19:25:43 +0100 wenzelm removed checkpoint interface;
Thu, 12 Oct 2006 22:57:45 +0200 wenzelm renamed enter_forward_proof to enter_proof_body;
Wed, 11 Oct 2006 22:55:22 +0200 wenzelm exit_local_theory: pass interactive flag;
Wed, 11 Oct 2006 00:27:38 +0200 wenzelm added type global_theory -- theory or local_theory;
Mon, 09 Oct 2006 19:37:07 +0200 wenzelm loop: disallow exit in secure mode;
Sat, 07 Oct 2006 01:31:19 +0200 wenzelm TheoryTarget;
Thu, 21 Sep 2006 19:04:20 +0200 wenzelm member (op =);
Wed, 09 Aug 2006 00:12:33 +0200 wenzelm global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
Fri, 14 Jul 2006 14:19:48 +0200 wenzelm keep/transaction: unified execution model (with debugging etc.);
Tue, 04 Jul 2006 18:39:59 +0200 wenzelm skip_proofs: do not skip proofs of schematic goals (warning);
Sun, 11 Jun 2006 21:59:17 +0200 wenzelm avoid unqualified exception;
Tue, 02 May 2006 20:42:41 +0200 wenzelm handle exception SYS_ERROR;
Fri, 17 Feb 2006 20:03:19 +0100 wenzelm checkpoint/copy_node: reset body context;
Wed, 15 Feb 2006 21:35:13 +0100 wenzelm added cases_node;
Mon, 06 Feb 2006 20:59:55 +0100 wenzelm added local_theory, with optional locale xname;
Fri, 27 Jan 2006 19:03:17 +0100 wenzelm swapped theory_context;
Thu, 19 Jan 2006 21:22:23 +0100 wenzelm keep: disable Output.toplevel_errors;
Sat, 14 Jan 2006 17:14:18 +0100 wenzelm sane ERROR vs. TOPLEVEL_ERROR handling;
Fri, 13 Jan 2006 01:13:00 +0100 wenzelm removed obsolete sign_of;
Tue, 10 Jan 2006 19:33:41 +0100 wenzelm added context_of -- generic context;
Fri, 06 Jan 2006 15:18:21 +0100 wenzelm transactions now always see quasi-functional intermediate checkpoint;
Thu, 05 Jan 2006 22:30:00 +0100 wenzelm hide type datatype node;
Wed, 04 Jan 2006 00:52:47 +0100 wenzelm added theory_to_theory_to_proof, which may change theory before entering the proof;
Wed, 04 Jan 2006 00:52:45 +0100 wenzelm tuned;
Tue, 18 Oct 2005 17:59:37 +0200 wenzelm simplified interfaces proof/proof' etc.: perform ProofHistory.apply(s)/current internally;
Tue, 20 Sep 2005 14:03:42 +0200 wenzelm get_interrupt: special handling of IO.io now in ML-Systems/smlnj-basis-compat.ML;
Tue, 20 Sep 2005 13:17:32 +0200 paulson uniform handling of interrupts
Sat, 17 Sep 2005 12:18:08 +0200 wenzelm theory_to_proof: check theory of initial proof state, which must not be changed;
Fri, 16 Sep 2005 11:38:49 +0200 paulson catching exception Io
Tue, 13 Sep 2005 22:19:48 +0200 wenzelm added three_buffersN, print3;
Mon, 12 Sep 2005 12:11:17 +0200 wenzelm added interact flag to control mode of excursions;
Sun, 11 Sep 2005 20:02:51 +0200 wenzelm excursion: interactive if debug;
Mon, 05 Sep 2005 17:38:24 +0200 wenzelm added assert, command;
Wed, 31 Aug 2005 15:46:47 +0200 wenzelm added no_body_context;
Thu, 18 Aug 2005 11:17:50 +0200 wenzelm proof_to_theory_context: interaction flag;
Tue, 16 Aug 2005 13:42:47 +0200 wenzelm state: body context;
Wed, 13 Jul 2005 16:07:37 +0200 wenzelm added print_state_hook;
Wed, 06 Jul 2005 20:00:41 +0200 wenzelm debug: exception_trace;
Mon, 04 Jul 2005 17:08:19 +0200 wenzelm tuned;
Fri, 01 Jul 2005 14:42:02 +0200 wenzelm added profile flag;
Wed, 29 Jun 2005 15:13:37 +0200 wenzelm added print': print depending on print_mode;
Mon, 20 Jun 2005 22:14:04 +0200 wenzelm tuned;
Fri, 17 Jun 2005 18:33:37 +0200 wenzelm Context.DATA_FAIL;
Thu, 02 Jun 2005 18:29:58 +0200 wenzelm tuned msgs;
Mon, 23 May 2005 14:56:36 +0200 wenzelm node_trans: revert to original transaction code (pre 1.54);
Tue, 17 May 2005 10:19:44 +0200 wenzelm tuned;
Thu, 07 Apr 2005 09:27:09 +0200 wenzelm improved exn_message;
Sat, 26 Mar 2005 18:20:29 +0100 paulson new display of theory stamps
Thu, 10 Mar 2005 17:48:36 +0100 ballarin Registrations of global locale interpretations: improved, better naming.
Wed, 09 Mar 2005 18:44:52 +0100 ballarin First version of global registration command.
Thu, 03 Mar 2005 12:43:01 +0100 skalberg Move towards standard functions.
Wed, 23 Feb 2005 15:19:00 +0100 berghofe Modified node_trans to avoid duplication of signature stamps
Sun, 13 Feb 2005 17:15:14 +0100 skalberg Deleted Library.option type.
Thu, 10 Feb 2005 12:06:40 +0100 ballarin Toplevel.debug for debugging in Isar.
Tue, 11 Jan 2005 14:14:39 +0100 berghofe excursion_result now also passes previous state to presentation functions.
Mon, 11 Oct 2004 10:51:19 +0200 berghofe Some changes to allow skipping of proof scripts.
Mon, 21 Jun 2004 16:39:58 +0200 wenzelm added >>> : transition list -> unit;
Mon, 21 Jun 2004 10:25:57 +0200 kleing Merged in license change from Isabelle2004
Sat, 12 Jun 2004 22:45:35 +0200 wenzelm added name_of, source_of, source;
Sat, 29 May 2004 15:01:36 +0200 wenzelm Output.timing;
Mon, 07 Jul 2003 17:58:21 +0200 nipkow A patch by david aspinall
Thu, 08 Aug 2002 23:51:24 +0200 wenzelm exception SIMPROC_FAIL: solid error reporting of simprocs;
Thu, 28 Feb 2002 21:31:47 +0100 wenzelm use ignore_interrupt, raise_interrupt;
Tue, 12 Feb 2002 20:34:02 +0100 wenzelm ANTIQUOTE_FAIL;
Sat, 08 Dec 2001 14:43:48 +0100 wenzelm tuned print_state interfaces;
Wed, 28 Nov 2001 23:29:48 +0100 wenzelm added proof_to_theory';
Sun, 04 Nov 2001 21:00:28 +0100 wenzelm simplified Proof.init_state:
Wed, 31 Oct 2001 22:02:33 +0100 wenzelm Proof.init_state thy None;
Mon, 22 Oct 2001 18:04:11 +0200 wenzelm Display.current_goals_markers;
less more (0) -120 tip