src/HOL/Imperative_HOL/Heap_Monad.thy
Sun, 02 Nov 2014 18:21:45 +0100 wenzelm modernized header uniformly as section;
Thu, 02 Oct 2014 11:33:04 +0200 haftmann tuned Heap_Monad.successE
Thu, 18 Sep 2014 18:48:04 +0200 haftmann tuned data structure
Thu, 11 Sep 2014 19:32:36 +0200 blanchet updated news
Tue, 09 Sep 2014 20:51:36 +0200 blanchet use 'datatype_new' (soon to be renamed 'datatype') in Isabelle's libraries
Sat, 16 Aug 2014 18:08:55 +0200 wenzelm updated to named_theorems;
Mon, 30 Jun 2014 08:00:36 +0200 haftmann qualified String.explode and String.implode
Sun, 09 Feb 2014 21:37:27 +0100 haftmann dropped legacy finally
Sat, 25 Jan 2014 23:50:49 +0100 haftmann prefer explicit code symbol type over ad-hoc name mangling
Thu, 05 Dec 2013 09:20:32 +0100 Andreas Lochbihler restrict admissibility to non-empty chains to allow more syntax-directed proof rules
Mon, 02 Sep 2013 16:28:11 +0200 Andreas Lochbihler move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
Wed, 24 Jul 2013 17:15:59 +0200 krauss derive specialized version of full fixpoint induction (with admissibility)
Fri, 12 Jul 2013 16:19:05 +0200 wenzelm localized and modernized adhoc-overloading (patch by Christian Sternagel);
Sun, 23 Jun 2013 21:16:07 +0200 haftmann migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
Tue, 18 Jun 2013 15:35:53 +0200 lammich Added parantheses to code_type for heap monad
Fri, 22 Mar 2013 00:39:16 +0100 krauss added rudimentary induction rule for partial_function (heap)
Fri, 15 Feb 2013 08:31:31 +0100 haftmann two target language numeral types: integer and natural, as replacement for code_numeral;
Tue, 05 Jun 2012 07:10:51 +0200 haftmann prefer sys.error over plain error in Scala to avoid deprecation warning
Tue, 05 Jun 2012 07:05:56 +0200 haftmann prefer records with speaking labels over deeply nested tuples
Thu, 29 Dec 2011 10:47:55 +0100 haftmann attribute code_abbrev superseedes code_unfold_post; tuned names and spacing
Fri, 28 Oct 2011 23:41:16 +0200 wenzelm tuned Named_Thms: proper binding;
Fri, 21 Oct 2011 11:17:14 +0200 bulwahn replacing code_inline by code_unfold, removing obsolete code_unfold, code_inline del now that the ancient code generator is removed
Wed, 07 Sep 2011 13:51:37 +0200 bulwahn adapting Imperative HOL serializer to changes of the iterm datatype in the code generator
Fri, 12 Aug 2011 14:45:50 -0700 huffman make more HOL theories work with separate set type
Thu, 09 Jun 2011 16:34:49 +0200 wenzelm discontinued Name.variant to emphasize that this is old-style / indirect;
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, 29 Dec 2010 17:34:41 +0100 wenzelm explicit file specifications -- avoid secondary load path;
Mon, 22 Nov 2010 09:37:39 +0100 haftmann renamed slightly ambivalent crel to effect
Fri, 29 Oct 2010 14:03:02 +0200 haftmann tuned structure of theory
Fri, 29 Oct 2010 13:49:49 +0200 haftmann remove term_of equations for Heap type explicitly
Tue, 26 Oct 2010 12:19:02 +0200 krauss added Heap monad instance of partial_function package
Mon, 13 Sep 2010 11:13:15 +0200 nipkow renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
Thu, 09 Sep 2010 14:38:14 +0200 bulwahn changing String.literal to a type instead of a datatype
Tue, 07 Sep 2010 10:05:19 +0200 nipkow expand_fun_eq -> ext_iff
Wed, 01 Sep 2010 15:51:49 +0200 haftmann Graph.map, in analogy to Table.map
Wed, 01 Sep 2010 11:09:50 +0200 haftmann do not print object frame around Scala includes -- this is in the responsibility of the user
Thu, 26 Aug 2010 12:19:49 +0200 haftmann prevent line breaks after Scala symbolic operators
Thu, 26 Aug 2010 10:16:22 +0200 haftmann code_include Scala: qualify module nmae
Fri, 13 Aug 2010 14:45:07 +0200 haftmann lemma execute_bind_case
Thu, 12 Aug 2010 09:00:19 +0200 haftmann tuned
Thu, 12 Aug 2010 08:58:32 +0200 haftmann tuned
Thu, 29 Jul 2010 09:56:59 +0200 haftmann proper unit type in transformed program
Mon, 26 Jul 2010 11:09:44 +0200 haftmann use Natural as index type for Haskell and Scala
Fri, 23 Jul 2010 10:58:13 +0200 haftmann avoid unreliable Haskell Int type
Mon, 19 Jul 2010 11:55:42 +0200 haftmann dropped superfluous prefixes
Fri, 16 Jul 2010 15:28:22 +0200 haftmann first roughly working version of Imperative HOL for Scala
Fri, 16 Jul 2010 13:58:29 +0200 haftmann a first sketch for Imperative HOL witht Scala
Fri, 16 Jul 2010 10:23:21 +0200 haftmann fragments of Scala
Thu, 15 Jul 2010 08:14:05 +0200 haftmann dropped spurious export_code
Wed, 14 Jul 2010 17:15:58 +0200 haftmann repaired some implementations of imperative operations
Wed, 14 Jul 2010 16:45:30 +0200 haftmann avoid ambiguities; tuned
Wed, 14 Jul 2010 12:27:44 +0200 haftmann dropped M suffix; added predicate monad bind
Tue, 13 Jul 2010 12:00:11 +0200 krauss Heap_Monad uses Monad_Syntax
Tue, 13 Jul 2010 11:38:03 +0200 haftmann theorem collections do not contain default rules any longer
Mon, 12 Jul 2010 16:19:15 +0200 haftmann split off mrec into separate theory
Mon, 12 Jul 2010 16:05:08 +0200 haftmann spelt out relational framework in a consistent way
Fri, 09 Jul 2010 16:58:44 +0200 haftmann pervasive success combinator
Fri, 09 Jul 2010 10:08:10 +0200 haftmann avoid slightly odd "M" suffix; rename mapM to fold_map (fold_map_abort would be more correct, though)
Fri, 09 Jul 2010 09:48:53 +0200 haftmann guard combinator
less more (0) -60 tip