Fri, 10 Sep 2021 14:59:19 +0200 |
wenzelm |
clarified signature: more scalable operations;
|
file |
diff |
annotate
|
Mon, 06 Sep 2021 12:23:06 +0200 |
wenzelm |
clarified modules;
|
file |
diff |
annotate
|
Wed, 04 Aug 2021 19:41:59 +0200 |
wenzelm |
clarified export of formal entities: name space info is always present, but content depends on option "export_theory";
|
file |
diff |
annotate
|
Fri, 21 May 2021 12:29:29 +0200 |
wenzelm |
clarified modules;
|
file |
diff |
annotate
|
Mon, 12 Oct 2020 07:25:38 +0000 |
haftmann |
consolidated names and operations
|
file |
diff |
annotate
|
Fri, 14 Aug 2020 14:40:24 +0200 |
wenzelm |
clarified signature;
|
file |
diff |
annotate
|
Fri, 06 Dec 2019 17:05:52 +0100 |
wenzelm |
suppress record types: not working properly;
|
file |
diff |
annotate
|
Fri, 06 Dec 2019 15:53:09 +0100 |
wenzelm |
clarified signature;
|
file |
diff |
annotate
|
Fri, 06 Dec 2019 15:44:55 +0100 |
wenzelm |
export datatypes;
|
file |
diff |
annotate
|
Fri, 06 Dec 2019 11:43:29 +0100 |
wenzelm |
clarified modules;
|
file |
diff |
annotate
|
Mon, 02 Dec 2019 15:04:38 +0100 |
wenzelm |
proper spec_rule name via naming/binding/Morphism.binding;
|
file |
diff |
annotate
|
Fri, 29 Nov 2019 20:57:04 +0100 |
wenzelm |
more informative spec rules: optional name;
|
file |
diff |
annotate
|
Fri, 09 Aug 2019 17:14:49 +0200 |
wenzelm |
formal position for PThm nodes;
|
file |
diff |
annotate
|
Tue, 26 Mar 2019 22:13:36 +0100 |
wenzelm |
more informative Spec_Rules.Equational, notably primrec argument types;
|
file |
diff |
annotate
|
Sat, 05 Jan 2019 17:24:33 +0100 |
wenzelm |
isabelle update -u control_cartouches;
|
file |
diff |
annotate
|
Fri, 04 Jan 2019 23:22:53 +0100 |
wenzelm |
isabelle update -u control_cartouches;
|
file |
diff |
annotate
|
Thu, 11 Jan 2018 13:48:17 +0100 |
wenzelm |
uniform use of Standard ML op-infix -- eliminated warnings;
|
file |
diff |
annotate
|
Wed, 10 Jan 2018 15:25:09 +0100 |
nipkow |
ran isabelle update_op on all sources
|
file |
diff |
annotate
|
Sun, 02 Jul 2017 20:13:38 +0200 |
haftmann |
proper concept of code declaration wrt. atomicity and Isar declarations
|
file |
diff |
annotate
|
Sat, 29 Oct 2016 00:39:33 +0200 |
blanchet |
avoid code generator warnings when deleting equations in BNF commands (e.g., datatype) in locales with assumptions and similar
|
file |
diff |
annotate
|
Mon, 12 Sep 2016 16:51:55 +0200 |
blanchet |
avoid warning triggered by code generator
|
file |
diff |
annotate
|
Fri, 17 Jun 2016 21:25:59 +0200 |
blanchet |
avoid runtime warning with discriminators due to 'Code.del_eqn'
|
file |
diff |
annotate
|
Mon, 06 Jun 2016 21:28:46 +0200 |
haftmann |
explicit tagging of code equations de-baroquifies interface
|
file |
diff |
annotate
|
Sun, 29 May 2016 15:40:25 +0200 |
wenzelm |
clarified check_open_spec / read_open_spec;
|
file |
diff |
annotate
|
Thu, 28 Apr 2016 09:43:11 +0200 |
wenzelm |
support 'assumes' in specifications, e.g. 'definition', 'inductive';
|
file |
diff |
annotate
|
Wed, 17 Feb 2016 11:54:34 +0100 |
blanchet |
allow predicator instead of map function in 'primrec'
|
file |
diff |
annotate
|
Mon, 15 Feb 2016 18:27:17 +0100 |
blanchet |
tuning
|
file |
diff |
annotate
|
Mon, 15 Feb 2016 12:48:10 +0100 |
blanchet |
tuning
|
file |
diff |
annotate
|
Tue, 01 Dec 2015 13:07:40 +0100 |
blanchet |
tuned whitespace
|
file |
diff |
annotate
|
Mon, 02 Nov 2015 21:49:49 +0100 |
blanchet |
allow selectors and discriminators with same name as type
|
file |
diff |
annotate
|
Tue, 06 Oct 2015 12:01:07 +0200 |
traytel |
collect the names from goals in favor of fragile exports
|
file |
diff |
annotate
|
Thu, 01 Oct 2015 17:32:07 +0200 |
blanchet |
export '_cmd' functions
|
file |
diff |
annotate
|
Fri, 25 Sep 2015 23:01:34 +0200 |
traytel |
restructure fresh variable generation to make exports more wellformed
|
file |
diff |
annotate
|
Fri, 25 Sep 2015 23:01:31 +0200 |
traytel |
more canonical context threading
|
file |
diff |
annotate
|
Fri, 04 Sep 2015 21:40:59 +0200 |
wenzelm |
close derivation *before* splitting conjuncts, like Goal.prove_common (see also 757cad5a3fe9) -- potential improvement of performance;
|
file |
diff |
annotate
|
Thu, 03 Sep 2015 16:41:43 +0200 |
traytel |
use open/close_target rather than Local_Theory.restore to get polymorphic definitions;
|
file |
diff |
annotate
|
Tue, 28 Jul 2015 20:59:39 +0200 |
wenzelm |
more explicit context;
|
file |
diff |
annotate
|
Mon, 27 Jul 2015 17:44:55 +0200 |
wenzelm |
tuned signature;
|
file |
diff |
annotate
|
Thu, 16 Jul 2015 12:23:22 +0200 |
traytel |
{r,e,d,f}tac with proper context in BNF
|
file |
diff |
annotate
|
Sun, 05 Jul 2015 15:02:30 +0200 |
wenzelm |
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
|
file |
diff |
annotate
|
Sat, 09 May 2015 14:10:10 +0200 |
blanchet |
took out unreliable 'blast' from tactic altogether
|
file |
diff |
annotate
|
Sun, 03 May 2015 18:14:11 +0200 |
blanchet |
made split-rule tactic go beyond constructors with 20 arguments
|
file |
diff |
annotate
|
Mon, 06 Apr 2015 17:06:48 +0200 |
wenzelm |
@{command_spec} is superseded by @{command_keyword};
|
file |
diff |
annotate
|
Tue, 31 Mar 2015 00:11:54 +0200 |
wenzelm |
tuned signature;
|
file |
diff |
annotate
|
Sat, 07 Mar 2015 21:32:31 +0100 |
wenzelm |
clarified Drule.gen_all: observe context more carefully;
|
file |
diff |
annotate
|
Fri, 06 Mar 2015 15:58:56 +0100 |
wenzelm |
Thm.cterm_of and Thm.ctyp_of operate on local context;
|
file |
diff |
annotate
|
Wed, 04 Mar 2015 19:53:18 +0100 |
wenzelm |
tuned signature -- prefer qualified names;
|
file |
diff |
annotate
|
Tue, 03 Mar 2015 19:08:04 +0100 |
traytel |
eliminated some clones of Proof_Context.cterm_of
|
file |
diff |
annotate
|
Mon, 05 Jan 2015 21:24:14 +0100 |
blanchet |
generate [code] only with 'code' plugin enabled
|
file |
diff |
annotate
|
Mon, 05 Jan 2015 10:09:42 +0100 |
blanchet |
added plugins syntax to prim(co)rec
|
file |
diff |
annotate
|
Mon, 05 Jan 2015 06:56:15 +0100 |
blanchet |
tuning
|
file |
diff |
annotate
|
Fri, 19 Dec 2014 11:18:58 +0100 |
desharna |
generate 'disc_eq_case' for Ctr_Sugars
|
file |
diff |
annotate
|
Mon, 05 Jan 2015 06:56:15 +0100 |
blanchet |
tuning
|
file |
diff |
annotate
|
Fri, 19 Dec 2014 11:17:23 +0100 |
desharna |
generate 'case_distrib' for Ctr_Sugars
|
file |
diff |
annotate
|
Wed, 26 Nov 2014 20:05:34 +0100 |
wenzelm |
renamed "pairself" to "apply2", in accordance to @{apply 2};
|
file |
diff |
annotate
|
Thu, 13 Nov 2014 17:28:11 +0100 |
blanchet |
tuning
|
file |
diff |
annotate
|
Mon, 10 Nov 2014 21:49:48 +0100 |
wenzelm |
proper context for assume_tac (atac remains as fall-back without context);
|
file |
diff |
annotate
|
Wed, 15 Oct 2014 17:18:08 +0200 |
blanchet |
made SML/NJ happier
|
file |
diff |
annotate
|
Tue, 14 Oct 2014 16:17:34 +0200 |
desharna |
add 'kind' to 'cr_sugar'
|
file |
diff |
annotate
|
Mon, 13 Oct 2014 21:41:29 +0200 |
wenzelm |
tuned signature;
|
file |
diff |
annotate
|
Mon, 13 Oct 2014 19:34:10 +0200 |
wenzelm |
clarified load order;
|
file |
diff |
annotate
|
Mon, 13 Oct 2014 18:45:48 +0200 |
wenzelm |
Local_Interpretation is superseded by Plugin with formal Plugin_Name management, avoiding undeclared strings;
|
file |
diff |
annotate
|
Wed, 08 Oct 2014 17:09:07 +0200 |
wenzelm |
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
|
file |
diff |
annotate
|
Mon, 15 Sep 2014 10:49:07 +0200 |
blanchet |
generate 'code' attribute only if 'code' plugin is enabled
|
file |
diff |
annotate
|
Thu, 11 Sep 2014 20:01:29 +0200 |
blanchet |
tuned comment
|
file |
diff |
annotate
|
Tue, 09 Sep 2014 22:33:43 +0200 |
blanchet |
avoid exception
|
file |
diff |
annotate
|
Tue, 09 Sep 2014 20:51:36 +0200 |
blanchet |
nicer case names in the N2M case, similar to those generated by the old package (e.g. 'Cons_tree' instead of just 'Cons')
|
file |
diff |
annotate
|
Tue, 09 Sep 2014 20:51:36 +0200 |
blanchet |
generalized 'datatype' LaTeX antiquotation and added 'codatatype'
|
file |
diff |
annotate
|
Mon, 08 Sep 2014 19:21:07 +0200 |
blanchet |
honour sorts in N2M
|
file |
diff |
annotate
|
Mon, 08 Sep 2014 15:11:37 +0200 |
blanchet |
use right sort constraints
|
file |
diff |
annotate
|
Mon, 08 Sep 2014 14:03:46 +0200 |
blanchet |
made SML/NJ happire
|
file |
diff |
annotate
|
Fri, 05 Sep 2014 00:41:01 +0200 |
blanchet |
pretend code generation is a ctr_sugar plugin
|
file |
diff |
annotate
|
Fri, 05 Sep 2014 00:41:01 +0200 |
blanchet |
added 'plugins' option to control which hooks are enabled
|
file |
diff |
annotate
|
Fri, 05 Sep 2014 00:41:01 +0200 |
blanchet |
introduced mechanism to filter interpretations
|
file |
diff |
annotate
|
Fri, 05 Sep 2014 00:41:01 +0200 |
blanchet |
fixed infinite loops in 'register' functions + more uniform API
|
file |
diff |
annotate
|
Fri, 05 Sep 2014 00:41:01 +0200 |
blanchet |
named interpretations
|
file |
diff |
annotate
|
Fri, 05 Sep 2014 00:41:00 +0200 |
blanchet |
centralized and cleaned up naming handling
|
file |
diff |
annotate
|
Thu, 04 Sep 2014 11:20:59 +0200 |
blanchet |
tweaked setup for datatype realizer
|
file |
diff |
annotate
|
Wed, 03 Sep 2014 22:49:05 +0200 |
blanchet |
introduced local interpretation mechanism for BNFs, to solve issues with datatypes in locales
|
file |
diff |
annotate
|
Wed, 03 Sep 2014 22:47:48 +0200 |
blanchet |
tuned ctr_sugar database handling
|
file |
diff |
annotate
|
Wed, 03 Sep 2014 00:06:23 +0200 |
blanchet |
take out 'x = C' of the simplifier for unit types
|
file |
diff |
annotate
|
Mon, 01 Sep 2014 16:17:47 +0200 |
blanchet |
added theory-based getters for convenience
|
file |
diff |
annotate
|
Mon, 01 Sep 2014 16:17:47 +0200 |
blanchet |
made transfer functions slightly more general
|
file |
diff |
annotate
|
Mon, 18 Aug 2014 17:20:14 +0200 |
blanchet |
tuning
|
file |
diff |
annotate
|
Mon, 18 Aug 2014 17:20:13 +0200 |
blanchet |
added collection theorem for consistency and convenience
|
file |
diff |
annotate
|
Mon, 18 Aug 2014 17:19:58 +0200 |
blanchet |
reordered some (co)datatype property names for more consistency
|
file |
diff |
annotate
|
Sun, 10 Aug 2014 14:34:43 +0200 |
wenzelm |
merged -- with manual conflict resolution for src/HOL/SMT_Examples/SMT_Examples.certs2, src/HOL/SMT_Examples/SMT_Word_Examples.certs2, src/Doc/Prog_Prove/document/intro-isabelle.tex;
|
file |
diff |
annotate
|
Wed, 30 Jul 2014 09:40:28 +0200 |
blanchet |
correctly resolve selector names in default value equations
|
file |
diff |
annotate
|
Wed, 30 Jul 2014 10:50:28 +0200 |
desharna |
generate 'set_induct' theorem for codatatypes
|
file |
diff |
annotate
|
Thu, 24 Jul 2014 00:24:00 +0200 |
blanchet |
repaired named derivations
|
file |
diff |
annotate
|
Thu, 24 Jul 2014 00:24:00 +0200 |
blanchet |
use the noted theorem whenever possible, also in 'BNF_Def'
|
file |
diff |
annotate
|
Thu, 24 Jul 2014 00:24:00 +0200 |
blanchet |
use the noted theorem whenever possible, because it has a named derivation (leading to cleaner proof terms)
|
file |
diff |
annotate
|
Mon, 16 Jun 2014 19:18:10 +0200 |
blanchet |
fixed postprocessing of 'coinduct' formula to obtain right property format (without needless hypotheses)
|
file |
diff |
annotate
|
Tue, 10 Jun 2014 12:16:22 +0200 |
blanchet |
use 'where' clause for selector default value syntax
|
file |
diff |
annotate
|
Tue, 27 May 2014 17:32:42 +0200 |
blanchet |
don't generate discriminators and selectors for 'datatype_new' unless the user asked for it
|
file |
diff |
annotate
|
Mon, 26 May 2014 16:32:55 +0200 |
blanchet |
got rid of '=:' squiggly
|
file |
diff |
annotate
|
Mon, 26 May 2014 16:32:51 +0200 |
blanchet |
use '%x. x = C' as default discriminator for nullary constructor C, instead of relying on odd '=:' syntax
|
file |
diff |
annotate
|
Wed, 21 May 2014 14:09:42 +0200 |
blanchet |
move exhaust first, for technical reasons
|
file |
diff |
annotate
|
Mon, 05 May 2014 09:30:20 +0200 |
blanchet |
simplify selectors in code views
|
file |
diff |
annotate
|
Mon, 28 Apr 2014 00:54:31 +0200 |
blanchet |
restored naming trick
|
file |
diff |
annotate
|
Wed, 23 Apr 2014 11:29:39 +0200 |
blanchet |
made SML/NJ happier
|
file |
diff |
annotate
|
Thu, 10 Apr 2014 17:48:17 +0200 |
kuncar |
revert idiomatic handling of namings from 5a93b8f928a2 because in combination with Named_Target.theory_init global names are sometimes created
|
file |
diff |
annotate
|
Thu, 10 Apr 2014 17:48:16 +0200 |
kuncar |
don't forget to init Interpretation and transfer theorems in the interpretation hook
|
file |
diff |
annotate
|
Thu, 03 Apr 2014 10:51:24 +0200 |
blanchet |
added same idiomatic handling of namings for Ctr_Sugar/BNF-related interpretation hooks as for typedef and (old-style) datatypes
|
file |
diff |
annotate
|
Tue, 01 Apr 2014 10:51:29 +0200 |
blanchet |
added 'ctr_sugar' interpretation hook
|
file |
diff |
annotate
|
Mon, 17 Feb 2014 13:31:42 +0100 |
blanchet |
name derivations in 'primrec' for code extraction from proof terms
|
file |
diff |
annotate
|
Fri, 14 Feb 2014 15:03:24 +0100 |
blanchet |
allow different functions to recurse on the same type, like in the old package
|
file |
diff |
annotate
|
Fri, 14 Feb 2014 07:53:46 +0100 |
blanchet |
more precise spec rules for selectors
|
file |
diff |
annotate
|
Fri, 14 Feb 2014 07:53:46 +0100 |
blanchet |
removed needless robustness (no longer needed thanks to new syntax)
|
file |
diff |
annotate
|
Fri, 14 Feb 2014 07:53:46 +0100 |
blanchet |
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
|
file |
diff |
annotate
|
Fri, 14 Feb 2014 07:53:46 +0100 |
blanchet |
renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
|
file |
diff |
annotate
|
Fri, 14 Feb 2014 07:53:45 +0100 |
blanchet |
have 'Ctr_Sugar' register its 'Spec_Rules'
|
file |
diff |
annotate
|
Wed, 12 Feb 2014 17:36:00 +0100 |
blanchet |
iteration n in the 'default' vs. 'update_new' vs. 'update' saga -- 'update' makes sense now that we honor the canonical order on 'merge' (as opposed to raising 'DUP')
|
file |
diff |
annotate
|
Wed, 12 Feb 2014 08:35:56 +0100 |
blanchet |
killed 'rep_compat' option
|
file |
diff |
annotate
|
Wed, 12 Feb 2014 08:35:56 +0100 |
blanchet |
made 'ctr_sugar' more friendly to the 'datatype_realizer'
|
file |
diff |
annotate
|
Wed, 12 Feb 2014 08:35:56 +0100 |
blanchet |
use right local theory -- shows up when 'no_discs_sels' is set
|
file |
diff |
annotate
|
Wed, 12 Feb 2014 08:35:56 +0100 |
blanchet |
more liberal merging of BNFs and constructor sugar
|
file |
diff |
annotate
|
Fri, 07 Feb 2014 10:44:04 +0100 |
blanchet |
reverted a87e49f4336d -- overwriting of data entries yields to merge problems later
|
file |
diff |
annotate
|
Thu, 06 Feb 2014 17:05:47 +0100 |
blanchet |
allow multiple registration of the same type, the last wins
|
file |
diff |
annotate
|
Thu, 06 Feb 2014 00:03:12 +0100 |
blanchet |
tuning
|
file |
diff |
annotate
|