src/HOL/Tools/Function/partial_function.ML
Tue, 04 Jun 2019 15:14:19 +0200 wenzelm proper context;
Tue, 26 Mar 2019 22:13:36 +0100 wenzelm more informative Spec_Rules.Equational, notably primrec argument types;
Thu, 21 Feb 2019 09:15:07 +0000 haftmann streamlined specification interfaces
Fri, 04 Jan 2019 23:22:53 +0100 wenzelm isabelle update -u control_cartouches;
Mon, 19 Feb 2018 14:49:11 +0100 wenzelm tuned signature;
Wed, 06 Dec 2017 20:43:09 +0100 wenzelm prefer control symbol antiquotations;
Mon, 11 Sep 2017 18:36:13 +0200 wenzelm clarified signature: proper result;
Sat, 11 Jun 2016 16:41:11 +0200 wenzelm clarified syntax;
Mon, 30 May 2016 14:15:44 +0200 wenzelm allow 'for' fixes for multi_specs;
Fri, 27 May 2016 20:23:55 +0200 wenzelm tuned proofs, to allow unfold_abs_def;
Thu, 28 Apr 2016 09:43:11 +0200 wenzelm support 'assumes' in specifications, e.g. 'definition', 'inductive';
Mon, 18 Apr 2016 11:02:07 +0200 wenzelm avoid clash with function called "x";
Sun, 17 Apr 2016 22:10:09 +0200 wenzelm clarified bindings;
Sun, 17 Apr 2016 20:11:02 +0200 wenzelm clarified signature;
Sun, 17 Apr 2016 12:40:48 +0200 wenzelm clarified reported positions;
Sun, 17 Apr 2016 12:26:22 +0200 wenzelm operate on proper binding;
Wed, 13 Apr 2016 18:01:05 +0200 wenzelm eliminated "xname" and variants;
Thu, 07 Jan 2016 15:53:39 +0100 wenzelm more uniform treatment of package internals;
Mon, 14 Dec 2015 11:20:31 +0100 wenzelm tuned signature;
Sun, 13 Dec 2015 21:56:15 +0100 wenzelm more general types Proof.method / context_tactic;
Tue, 13 Oct 2015 09:21:15 +0200 haftmann prod_case as canonical name for product type eliminator
Sun, 06 Sep 2015 21:55:13 +0200 wenzelm do not expose low-level "_def" facts of 'function' definitions, to avoid potential confusion with the situation of plain 'definition';
Fri, 04 Sep 2015 14:00:13 +0200 wenzelm trim context for persistent storage;
Sun, 26 Jul 2015 17:24:54 +0200 wenzelm updated to infer_instantiate;
Wed, 08 Jul 2015 19:28:43 +0200 wenzelm Variable.focus etc.: optional bindings provided by user;
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);
Mon, 06 Apr 2015 17:06:48 +0200 wenzelm @{command_spec} is superseded by @{command_keyword};
Tue, 31 Mar 2015 00:11:54 +0200 wenzelm tuned signature;
Fri, 06 Mar 2015 15:58:56 +0100 wenzelm Thm.cterm_of and Thm.ctyp_of operate on local context;
Wed, 04 Mar 2015 19:53:18 +0100 wenzelm tuned signature -- prefer qualified names;
Tue, 03 Mar 2015 19:08:04 +0100 traytel eliminated some clones of Proof_Context.cterm_of
Tue, 10 Feb 2015 14:48:26 +0100 wenzelm proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
Thu, 30 Oct 2014 22:45:19 +0100 wenzelm eliminated aliases;
Sat, 16 Aug 2014 19:20:11 +0200 wenzelm updated to named_theorems;
Tue, 31 Dec 2013 14:29:16 +0100 wenzelm proper context for norm_hhf and derived operations;
Sat, 14 Dec 2013 17:28:05 +0100 wenzelm proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
Thu, 05 Dec 2013 09:20:32 +0100 Andreas Lochbihler restrict admissibility to non-empty chains to allow more syntax-directed proof rules
Tue, 12 Nov 2013 13:47:24 +0100 blanchet ported 'partial_function' to 'Ctr_Sugar' abstraction
Wed, 24 Jul 2013 17:15:59 +0200 krauss derive specialized version of full fixpoint induction (with admissibility)
Wed, 24 Jul 2013 15:29:23 +0200 krauss export mono_thm
Wed, 03 Jul 2013 00:08:29 +0200 krauss more robust instantiation: extract vars from conclusion instead of relying on fixed overall ordering
Mon, 20 May 2013 17:14:39 +0200 wenzelm more precise treatment of theory vs. Proof.context;
Thu, 18 Apr 2013 17:07:01 +0200 wenzelm simplifier uses proper Proof.context instead of historic type simpset;
Fri, 22 Mar 2013 00:39:14 +0100 krauss allow induction predicates with arbitrary arity (not just binary)
Fri, 16 Mar 2012 18:20:12 +0100 wenzelm outer syntax command definitions based on formal command_spec derived from theory header declarations;
Thu, 15 Mar 2012 20:07:00 +0100 wenzelm prefer formally checked @{keyword} parser;
Tue, 08 Nov 2011 11:44:37 +0100 wenzelm tuned;
Fri, 28 Oct 2011 23:41:16 +0200 wenzelm tuned Named_Thms: proper binding;
Tue, 20 Sep 2011 01:32:04 +0200 krauss match types when applying mono_thm -- previous export generalizes type variables;
Tue, 31 May 2011 11:16:52 +0200 krauss generate raw induction rule as instance of generic rule with careful treatment of currying
Mon, 23 May 2011 21:34:37 +0200 krauss also manage induction rule;
Mon, 23 May 2011 10:58:21 +0200 krauss separate initializations for different modes of partial_function -- generation of induction rules will be non-uniform
Wed, 27 Apr 2011 21:50:04 +0200 wenzelm clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
Mon, 18 Apr 2011 13:52:23 +0200 wenzelm standardized aliases of operations on tsig;
Sat, 16 Apr 2011 16:15:37 +0200 wenzelm modernized structure Proof_Context;
Thu, 24 Mar 2011 16:56:19 +0100 wenzelm added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
Sat, 08 Jan 2011 17:14:48 +0100 wenzelm misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
Mon, 13 Dec 2010 10:15:27 +0100 krauss eliminated dest_all_all_ctx
Tue, 26 Oct 2010 14:11:34 +0200 haftmann partial_function is a declaration command
Tue, 26 Oct 2010 13:19:31 +0200 krauss fixed confusion introduced in 008dc2d2c395
less more (0) -60 tip