src/Pure/Pure.thy
Tue, 06 Sep 2016 13:26:14 +0200 wenzelm strictly sequential abbrevs;
Tue, 02 Aug 2016 17:35:18 +0200 wenzelm support 'abbrevs' within theory header;
Sat, 16 Jul 2016 00:38:33 +0200 wenzelm information about proof outline with cases (sendback);
Fri, 15 Jul 2016 23:46:28 +0200 wenzelm singleton result for 'proof' command (without backtracking), e.g. relevant for well-defined output;
Mon, 11 Jul 2016 20:58:00 +0200 wenzelm clarified indentation;
Mon, 11 Jul 2016 16:36:29 +0200 wenzelm explicit kind "before_command";
Mon, 11 Jul 2016 10:43:54 +0200 wenzelm more indentation for quasi_command keywords;
Thu, 23 Jun 2016 11:01:14 +0200 wenzelm tuned signature;
Tue, 21 Jun 2016 17:35:45 +0200 wenzelm tuned;
Sat, 11 Jun 2016 16:41:11 +0200 wenzelm clarified syntax;
Fri, 10 Jun 2016 22:47:25 +0200 wenzelm added command 'unbundle';
Fri, 10 Jun 2016 12:45:34 +0200 wenzelm prefer hybrid 'bundle' command;
Thu, 09 Jun 2016 17:13:52 +0200 wenzelm clarified;
Thu, 09 Jun 2016 15:41:49 +0200 wenzelm support for bundle definition via target;
Sun, 29 May 2016 15:40:25 +0200 wenzelm clarified check_open_spec / read_open_spec;
Sat, 28 May 2016 21:38:58 +0200 wenzelm clarified 'axiomatization';
Sat, 14 May 2016 19:49:10 +0200 wenzelm toplevel theorem statements support 'if'/'for' eigen-context;
Thu, 28 Apr 2016 09:43:11 +0200 wenzelm support 'assumes' in specifications, e.g. 'definition', 'inductive';
Tue, 26 Apr 2016 22:39:17 +0200 wenzelm 'obtain' supports structured statements (similar to 'define');
Mon, 25 Apr 2016 19:41:39 +0200 wenzelm old 'def' is legacy;
Sun, 24 Apr 2016 21:31:14 +0200 wenzelm added Isar command 'define';
Wed, 13 Apr 2016 18:01:05 +0200 wenzelm eliminated "xname" and variants;
Sun, 10 Apr 2016 21:46:12 +0200 wenzelm more standard session build process, including browser_info;
Fri, 08 Apr 2016 20:15:20 +0200 wenzelm eliminated unused simproc identifier;
Thu, 07 Apr 2016 16:53:43 +0200 wenzelm more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
Thu, 07 Apr 2016 12:13:11 +0200 wenzelm Pure attribute setup is back to Pure/Isar/attrib.ML, where it can be editing continuously (see also 7eb0c04e4c40);
Tue, 05 Apr 2016 18:20:25 +0200 wenzelm support bootstrap from fresh SML environment, with syntax of Isabelle/ML or SML;
Tue, 05 Apr 2016 15:53:48 +0200 wenzelm tuned;
Mon, 04 Apr 2016 23:58:48 +0200 wenzelm more uniform ML file commands;
Mon, 04 Apr 2016 23:08:43 +0200 wenzelm tuned whitespace;
Mon, 04 Apr 2016 22:13:08 +0200 wenzelm tuned -- more explicit sections;
Mon, 04 Apr 2016 17:25:53 +0200 wenzelm clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
Mon, 04 Apr 2016 17:02:34 +0200 wenzelm clarified bootstrap -- more uniform use of ML files;
Tue, 01 Mar 2016 19:42:59 +0100 wenzelm ML debugger support in Pure (again, see 3565c9f407ec);
Sun, 14 Feb 2016 16:30:27 +0100 wenzelm command '\<proof>' is an alias for 'sorry', with different typesetting;
Wed, 13 Jan 2016 16:55:56 +0100 wenzelm removed old 'defs' command;
Sat, 19 Dec 2015 11:05:04 +0100 haftmann abandoned attempt to unify sublocale and interpretation into global theories
Wed, 16 Dec 2015 16:31:36 +0100 wenzelm rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
Sat, 14 Nov 2015 08:45:52 +0100 haftmann prefer "rewrites" and "defines" to note rewrite morphisms
Sat, 14 Nov 2015 08:45:52 +0100 haftmann coalesce permanent_interpretation.ML with interpretation.ML
Tue, 10 Nov 2015 20:10:17 +0100 wenzelm clarified modules;
Thu, 05 Nov 2015 00:02:30 +0100 wenzelm symbolic syntax "\<comment> text";
Wed, 04 Nov 2015 08:13:52 +0100 ballarin Keyword 'rewrites' identifies rewrite morphisms.
Tue, 06 Oct 2015 15:39:00 +0200 wenzelm added 'proposition' command;
Tue, 06 Oct 2015 15:14:28 +0200 wenzelm fewer aliases for toplevel theorem statements;
Tue, 22 Sep 2015 18:56:25 +0200 wenzelm separate command 'print_definitions';
Mon, 17 Aug 2015 19:34:15 +0200 wenzelm support for ML files with/without debugger information;
Fri, 17 Jul 2015 21:40:47 +0200 wenzelm skeleton for interactive debugger;
Wed, 08 Jul 2015 15:37:32 +0200 wenzelm clarified text folds: proof ... qed counts as extra block;
Wed, 01 Jul 2015 21:48:46 +0200 wenzelm clarified keyword categories;
Wed, 01 Jul 2015 21:29:57 +0200 wenzelm support for subgoal focus command;
Mon, 22 Jun 2015 20:36:33 +0200 wenzelm support 'when' statement, which corresponds to 'presume';
Thu, 11 Jun 2015 22:47:53 +0200 wenzelm support for 'consider' command;
Fri, 05 Jun 2015 13:41:06 +0200 wenzelm added Isar command 'supply';
Sat, 30 May 2015 22:18:12 +0200 wenzelm unused;
Thu, 16 Apr 2015 13:48:10 +0200 wenzelm clarified thy_deps;
Thu, 09 Apr 2015 20:42:32 +0200 wenzelm clarified keyword 'qualified' in accordance to a similar keyword from Haskell (despite unrelated Binding.qualified in Isabelle/ML);
Mon, 06 Apr 2015 22:11:01 +0200 wenzelm support for 'restricted' modifier: only qualified accesses outside the local scope;
Sat, 04 Apr 2015 14:04:11 +0200 wenzelm support private scope for individual local theory commands;
Wed, 01 Apr 2015 22:08:06 +0200 wenzelm added command 'experiment';
Thu, 13 Nov 2014 23:45:15 +0100 wenzelm uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
Fri, 07 Nov 2014 16:36:55 +0100 wenzelm plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
Thu, 06 Nov 2014 11:44:41 +0100 wenzelm simplified keyword kinds;
Sun, 02 Nov 2014 15:27:37 +0100 wenzelm uniform heading commands work in any context, even in theory header;
Sat, 01 Nov 2014 14:20:38 +0100 wenzelm eliminated spurious semicolons;
Fri, 31 Oct 2014 16:03:45 +0100 wenzelm discontinued Isar TTY loop;
Fri, 31 Oct 2014 15:15:10 +0100 wenzelm removed obsolete Proof General commands;
Fri, 31 Oct 2014 11:18:17 +0100 wenzelm discontinued Proof General;
Tue, 28 Oct 2014 11:42:51 +0100 wenzelm explicit keyword category for commands that may start a block;
Mon, 13 Oct 2014 19:34:10 +0200 wenzelm clarified load order;
Mon, 13 Oct 2014 15:45:23 +0200 wenzelm support for named plugins for definitional packages;
Tue, 07 Oct 2014 20:34:17 +0200 wenzelm more symbols;
Tue, 07 Oct 2014 20:27:31 +0200 wenzelm more cartouches;
Sun, 05 Oct 2014 16:05:17 +0200 wenzelm bibtex support in ML: document antiquotation @{cite} with markup;
Sun, 07 Sep 2014 17:51:28 +0200 haftmann separated class_deps command into separate file
Thu, 14 Aug 2014 10:48:40 +0200 wenzelm tuned signature -- prefer self-contained user-space tool;
Sun, 10 Aug 2014 16:13:12 +0200 wenzelm support for named collections of theorems in canonical order;
Fri, 04 Jul 2014 15:46:13 +0200 wenzelm suppress completion of obscure keyword;
Mon, 30 Jun 2014 10:34:28 +0200 wenzelm removed obsolete Isar system commands -- raw ML console is normally used for system tinkering;
Fri, 27 Jun 2014 16:04:56 +0200 wenzelm command 'print_term_bindings' supersedes 'print_binds';
Mon, 05 May 2014 15:17:07 +0200 wenzelm support print operations as asynchronous query;
Tue, 29 Apr 2014 22:52:15 +0200 wenzelm suppress slightly odd completions of "real";
Sat, 19 Apr 2014 17:23:05 +0200 wenzelm added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
Tue, 25 Mar 2014 13:18:10 +0100 wenzelm added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
Tue, 18 Mar 2014 16:16:28 +0100 wenzelm clarified modules;
Wed, 12 Mar 2014 21:58:48 +0100 wenzelm simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
Wed, 26 Feb 2014 10:45:06 +0100 wenzelm suppress completion of obscure keyword, avoid confusion with plain "simp";
Sun, 16 Feb 2014 17:25:03 +0100 wenzelm prefer user-space tool within Pure.thy;
Mon, 10 Feb 2014 22:08:18 +0100 wenzelm discontinued axiomatic 'classes', 'classrel', 'arities';
Sun, 26 Jan 2014 14:01:19 +0100 wenzelm discontinued obsolete attribute "standard";
Sat, 25 Jan 2014 18:34:05 +0100 wenzelm prefer self-contained user-space tool;
Sat, 25 Jan 2014 18:18:03 +0100 wenzelm define basic attributes in user-space Pure.thy -- which provides better hyperlinks and may serve as example;
Fri, 17 Jan 2014 20:31:39 +0100 wenzelm prefer user-space tool within Pure.thy;
Thu, 12 Dec 2013 21:28:13 +0100 wenzelm skeleton for Simplifier trace by Lars Hupel;
Wed, 18 Sep 2013 11:08:28 +0200 wenzelm moved module into plain Isabelle/ML user space;
Wed, 11 Sep 2013 20:16:28 +0200 wenzelm more explicit indication of 'done' as proof script element;
Mon, 02 Sep 2013 16:10:26 +0200 wenzelm more explicit indication of 'guess' as improper Isar (aka "script") element;
Sun, 07 Jul 2013 18:34:29 +0200 wenzelm discontinued command 'print_drafts';
Tue, 25 Jun 2013 20:35:50 +0200 wenzelm recover command tags from 4cf3f6153eb8 -- important for document preparation;
Mon, 24 Jun 2013 23:33:14 +0200 wenzelm improved "isabelle keywords" and "isabelle update_keywords" based on Isabelle/Scala, without requiring to build sessions first;
Mon, 24 Jun 2013 17:17:17 +0200 wenzelm back to keyword 'pr' :: diag as required for ProofGeneral command line -- reject this TTY command in Isabelle/jEdit by other means;
Mon, 24 Jun 2013 17:03:53 +0200 wenzelm more formal ProofGeneral command setup within theory Pure;
Sun, 23 Jun 2013 21:23:36 +0200 wenzelm proper diagnostic command 'print_state';
Sat, 25 May 2013 15:37:53 +0200 wenzelm syntax translations always depend on context;
Fri, 17 May 2013 20:53:28 +0200 wenzelm renamed 'print_configs' to 'print_options';
Wed, 15 May 2013 20:34:42 +0200 wenzelm moved files;
Wed, 15 May 2013 20:22:46 +0200 wenzelm maintain ProofGeneral preferences within ProofGeneral module;
Sat, 30 Mar 2013 14:57:06 +0100 wenzelm added 'print_defn_rules' command;
Thu, 28 Feb 2013 16:38:17 +0100 wenzelm discontinued obsolete 'axioms' command;
Wed, 27 Feb 2013 17:32:17 +0100 wenzelm discontinued redundant 'use' command;
Wed, 27 Feb 2013 12:45:19 +0100 wenzelm discontinued obsolete 'uses' within theory header;
Mon, 25 Feb 2013 13:29:19 +0100 wenzelm discontinued pointless command category "thy_schematic_goal" -- this is checked dynamically;
Mon, 25 Feb 2013 11:07:02 +0100 wenzelm reconsider 'pretty_setmargin' as "control" command (instead of "diag") -- it is stateful and Proof General legacy;
Tue, 19 Feb 2013 21:44:37 +0100 wenzelm back to non-schematic 'sublocale' and 'interpretation' (despite df8fc0567a3d) for more potential parallelism;
Wed, 19 Dec 2012 22:26:26 +0100 nipkow removed odd associativity of ==
Mon, 19 Nov 2012 22:34:17 +0100 wenzelm alternative completion for outer syntax keywords;
Wed, 26 Sep 2012 14:23:33 +0200 wenzelm tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
Sun, 26 Aug 2012 22:23:10 +0200 wenzelm entity markup for theory Pure, to enable hyperlinks etc.;
Wed, 22 Aug 2012 22:55:41 +0200 wenzelm prefer ML_file over old uses;
Thu, 02 Aug 2012 12:36:54 +0200 wenzelm more official command specifications, including source position;
less more (0) -120 tip