NEWS
author wenzelm
Mon, 23 Apr 2012 21:53:43 +0200
changeset 48576 5f9ce06f281e
parent 48565 05663f75964c
child 48577 3eef88e8496b
permissions -rw-r--r--
typedef with implicit set definition is considered legacy;
     1 Isabelle NEWS -- history user-relevant changes
     2 ==============================================
     3 
     4 New in Isabelle2012 (May 2012)
     5 ------------------------------
     6 
     7 *** General ***
     8 
     9 * Prover IDE (PIDE) improvements:
    10 
    11   - more robust Sledgehammer integration (as before the sledgehammer
    12     command line needs to be typed into the source buffer)
    13   - markup for bound variables
    14   - markup for types of term variables (e.g. displayed as tooltips)
    15   - support for user-defined Isar commands within the running session
    16   - improved support for Unicode outside original 16bit range
    17     e.g. glyph for \<A> (thanks to jEdit 4.5.1)
    18 
    19 * Updated and extended reference manuals ("isar-ref" and
    20 "implementation"); reduced remaining material in old "ref" manual.
    21 
    22 * Rule attributes in local theory declarations (e.g. locale or class)
    23 are now statically evaluated: the resulting theorem is stored instead
    24 of the original expression.  INCOMPATIBILITY in rare situations, where
    25 the historic accident of dynamic re-evaluation in interpretations
    26 etc. was exploited.
    27 
    28 * Commands 'lemmas' and 'theorems' allow local variables using 'for'
    29 declaration, and results are standardized before being stored.  Thus
    30 old-style "standard" after instantiation or composition of facts
    31 becomes obsolete.  Minor INCOMPATIBILITY, due to potential change of
    32 indices of schematic variables.
    33 
    34 * Renamed some inner syntax categories:
    35 
    36     num ~> num_token
    37     xnum ~> xnum_token
    38     xstr ~> str_token
    39 
    40 Minor INCOMPATIBILITY.  Note that in practice "num_const" or
    41 "num_position" etc. are mainly used instead (which also include
    42 position information via constraints).
    43 
    44 * Simplified configuration options for syntax ambiguity: see
    45 "syntax_ambiguity_warning" and "syntax_ambiguity_limit" in isar-ref
    46 manual.  Minor INCOMPATIBILITY.
    47 
    48 * Forward declaration of outer syntax keywords within the theory
    49 header -- minor INCOMPATIBILITY for user-defined commands.  Allow new
    50 commands to be used in the same theory where defined.
    51 
    52 
    53 *** Pure ***
    54 
    55 * Auxiliary contexts indicate block structure for specifications with
    56 additional parameters and assumptions.  Such unnamed contexts may be
    57 nested within other targets, like 'theory', 'locale', 'class',
    58 'instantiation' etc.  Results from the local context are generalized
    59 accordingly and applied to the enclosing target context.  Example:
    60 
    61   context
    62     fixes x y z :: 'a
    63     assumes xy: "x = y" and yz: "y = z"
    64   begin
    65 
    66   lemma my_trans: "x = z" using xy yz by simp
    67 
    68   end
    69 
    70   thm my_trans
    71 
    72 The most basic application is to factor-out context elements of
    73 several fixes/assumes/shows theorem statements, e.g. see
    74 ~~/src/HOL/Isar_Examples/Group_Context.thy
    75 
    76 Any other local theory specification element works within the "context
    77 ... begin ... end" block as well.
    78 
    79 * Bundled declarations associate attributed fact expressions with a
    80 given name in the context.  These may be later included in other
    81 contexts.  This allows to manage context extensions casually, without
    82 the logical dependencies of locales and locale interpretation.
    83 
    84 See commands 'bundle', 'include', 'including' etc. in the isar-ref
    85 manual.
    86 
    87 * Rule composition via attribute "OF" (or ML functions OF/MRS) is more
    88 tolerant against multiple unifiers, as long as the final result is
    89 unique.  (As before, rules are composed in canonical right-to-left
    90 order to accommodate newly introduced premises.)
    91 
    92 * Command 'definition' no longer exports the foundational "raw_def"
    93 into the user context.  Minor INCOMPATIBILITY, may use the regular
    94 "def" result with attribute "abs_def" to imitate the old version.
    95 
    96 * Attribute "abs_def" turns an equation of the form "f x y == t" into
    97 "f == %x y. t", which ensures that "simp" or "unfold" steps always
    98 expand it.  This also works for object-logic equality.  (Formerly
    99 undocumented feature.)
   100 
   101 * Discontinued old "prems" fact, which used to refer to the accidental
   102 collection of foundational premises in the context (already marked as
   103 legacy since Isabelle2011).
   104 
   105 * Obsolete command 'types' has been discontinued.  Use 'type_synonym'
   106 instead.  INCOMPATIBILITY.
   107 
   108 * Old code generator for SML and its commands 'code_module',
   109 'code_library', 'consts_code', 'types_code' have been discontinued.
   110 Use commands of the generic code generator instead.  INCOMPATIBILITY.
   111 
   112 * Redundant attribute "code_inline" has been discontinued. Use
   113 "code_unfold" instead.  INCOMPATIBILITY.
   114 
   115 * Dropped attribute "code_unfold_post" in favor of the its dual
   116 "code_abbrev", which yields a common pattern in definitions like
   117 
   118   definition [code_abbrev]: "f = t"
   119 
   120 INCOMPATIBILITY.
   121 
   122 * Sort constraints are now propagated in simultaneous statements, just
   123 like type constraints.  INCOMPATIBILITY in rare situations, where
   124 distinct sorts used to be assigned accidentally.  For example:
   125 
   126   lemma "P (x::'a::foo)" and "Q (y::'a::bar)"  -- "now illegal"
   127 
   128   lemma "P (x::'a)" and "Q (y::'a::bar)"
   129     -- "now uniform 'a::bar instead of default sort for first occurrence (!)"
   130 
   131 
   132 *** HOL ***
   133 
   134 * New tutorial "Programming and Proving in Isabelle/HOL" ("prog-prove").
   135 It completely supercedes "A Tutorial Introduction to Structured Isar Proofs",
   136 which has been removed. It supercedes "Isabelle/HOL, A Proof Assistant
   137 for Higher-Order Logic" as the recommended beginners tutorial
   138 but does not cover all of the material of that old tutorial.
   139 
   140 * Discontinued old Tutorial on Isar ("isar-overview");
   141 
   142 * Type 'a set is now a proper type constructor (just as before
   143 Isabelle2008).  Definitions mem_def and Collect_def have disappeared.
   144 Non-trivial INCOMPATIBILITY.  For developments keeping predicates and
   145 sets separate, it is often sufficient to rephrase sets S accidentally
   146 used as predicates by "%x. x : S" and predicates P accidentally used
   147 as sets by "{x. P x}".  Corresponding proofs in a first step should be
   148 pruned from any tinkering with former theorems mem_def and Collect_def
   149 as far as possible.
   150 
   151 For developments which deliberately mixed predicates and sets, a
   152 planning step is necessary to determine what should become a predicate
   153 and what a set.  It can be helpful to carry out that step in
   154 Isabelle2011-1 before jumping right into the current release.
   155 
   156 * The representation of numerals has changed.  Datatype "num"
   157 represents strictly positive binary numerals, along with functions
   158 "numeral :: num => 'a" and "neg_numeral :: num => 'a" to represent
   159 positive and negated numeric literals, respectively. (See definitions
   160 in ~~/src/HOL/Num.thy.) Potential INCOMPATIBILITY, some user theories
   161 may require adaptations as follows:
   162 
   163   - Theorems with number_ring or number_semiring constraints: These
   164     classes are gone; use comm_ring_1 or comm_semiring_1 instead.
   165 
   166   - Theories defining numeric types: Remove number, number_semiring,
   167     and number_ring instances. Defer all theorems about numerals until
   168     after classes one and semigroup_add have been instantiated.
   169 
   170   - Numeral-only simp rules: Replace each rule having a "number_of v"
   171     pattern with two copies, one for numeral and one for neg_numeral.
   172 
   173   - Theorems about subclasses of semiring_1 or ring_1: These classes
   174     automatically support numerals now, so more simp rules and
   175     simprocs may now apply within the proof.
   176 
   177   - Definitions and theorems using old constructors Pls/Min/Bit0/Bit1:
   178     Redefine using other integer operations.
   179 
   180 * Code generation by default implements sets as container type rather
   181 than predicates.  INCOMPATIBILITY.
   182 
   183 * New proof import from HOL Light: Faster, simpler, and more scalable.
   184 Requires a proof bundle, which is available as an external component.
   185 Discontinued old (and mostly dead) Importer for HOL4 and HOL Light.
   186 INCOMPATIBILITY.
   187 
   188 * New type synonym 'a rel = ('a * 'a) set
   189 
   190 * Typedef with implicit set definition is considered legacy.  Use
   191 "typedef (open)" form instead, which will eventually become the
   192 default.
   193 
   194 * More default pred/set conversions on a couple of relation operations
   195 and predicates.  Added powers of predicate relations.  Consolidation
   196 of some relation theorems:
   197 
   198   converse_def ~> converse_unfold
   199   rel_comp_def ~> relcomp_unfold
   200   symp_def ~> (dropped, use symp_def and sym_def instead)
   201   transp_def ~> transp_trans
   202   Domain_def ~> Domain_unfold
   203   Range_def ~> Domain_converse [symmetric]
   204 
   205 Generalized theorems INF_INT_eq, INF_INT_eq2, SUP_UN_eq, SUP_UN_eq2.
   206 
   207 See theory "Relation" for examples for making use of pred/set
   208 conversions by means of attributes "to_set" and "to_pred".
   209 
   210 INCOMPATIBILITY.
   211 
   212 * Renamed facts about the power operation on relations, i.e., relpow
   213 to match the constant's name:
   214 
   215   rel_pow_1 ~> relpow_1
   216   rel_pow_0_I ~> relpow_0_I
   217   rel_pow_Suc_I ~> relpow_Suc_I
   218   rel_pow_Suc_I2 ~> relpow_Suc_I2
   219   rel_pow_0_E ~> relpow_0_E
   220   rel_pow_Suc_E ~> relpow_Suc_E
   221   rel_pow_E ~> relpow_E
   222   rel_pow_Suc_D2 ~> relpow_Suc_D2
   223   rel_pow_Suc_E2 ~> relpow_Suc_E2
   224   rel_pow_Suc_D2' ~> relpow_Suc_D2'
   225   rel_pow_E2 ~> relpow_E2
   226   rel_pow_add ~> relpow_add
   227   rel_pow_commute ~> relpow
   228   rel_pow_empty ~> relpow_empty:
   229   rtrancl_imp_UN_rel_pow ~> rtrancl_imp_UN_relpow
   230   rel_pow_imp_rtrancl ~> relpow_imp_rtrancl
   231   rtrancl_is_UN_rel_pow ~> rtrancl_is_UN_relpow
   232   rtrancl_imp_rel_pow ~> rtrancl_imp_relpow
   233   rel_pow_fun_conv ~> relpow_fun_conv
   234   rel_pow_finite_bounded1 ~> relpow_finite_bounded1
   235   rel_pow_finite_bounded ~> relpow_finite_bounded
   236   rtrancl_finite_eq_rel_pow ~> rtrancl_finite_eq_relpow
   237   trancl_finite_eq_rel_pow ~> trancl_finite_eq_relpow
   238   single_valued_rel_pow ~> single_valued_relpow
   239 
   240 INCOMPATIBILITY.
   241 
   242 * Theory Relation: Consolidated constant name for relation composition
   243 and corresponding theorem names:
   244 
   245   - Renamed constant rel_comp to relcomp.
   246 
   247   - Dropped abbreviation pred_comp. Use relcompp instead.
   248 
   249   - Renamed theorems:
   250 
   251     rel_compI ~> relcompI
   252     rel_compEpair ~> relcompEpair
   253     rel_compE ~> relcompE
   254     pred_comp_rel_comp_eq ~> relcompp_relcomp_eq
   255     rel_comp_empty1 ~> relcomp_empty1
   256     rel_comp_mono ~> relcomp_mono
   257     rel_comp_subset_Sigma ~> relcomp_subset_Sigma
   258     rel_comp_distrib ~> relcomp_distrib
   259     rel_comp_distrib2 ~> relcomp_distrib2
   260     rel_comp_UNION_distrib ~> relcomp_UNION_distrib
   261     rel_comp_UNION_distrib2 ~> relcomp_UNION_distrib2
   262     single_valued_rel_comp ~> single_valued_relcomp
   263     rel_comp_def ~> relcomp_unfold
   264     converse_rel_comp ~> converse_relcomp
   265     pred_compI ~> relcomppI
   266     pred_compE ~> relcomppE
   267     pred_comp_bot1 ~> relcompp_bot1
   268     pred_comp_bot2 ~> relcompp_bot2
   269     transp_pred_comp_less_eq ~> transp_relcompp_less_eq
   270     pred_comp_mono ~> relcompp_mono
   271     pred_comp_distrib ~> relcompp_distrib
   272     pred_comp_distrib2 ~> relcompp_distrib2
   273     converse_pred_comp ~> converse_relcompp
   274 
   275     finite_rel_comp ~> finite_relcomp
   276 
   277     set_rel_comp ~> set_relcomp
   278 
   279 INCOMPATIBILITY.
   280 
   281 * Theory Divides: Discontinued redundant theorems about div and mod.
   282 INCOMPATIBILITY, use the corresponding generic theorems instead.
   283 
   284   DIVISION_BY_ZERO ~> div_by_0, mod_by_0
   285   zdiv_self ~> div_self
   286   zmod_self ~> mod_self
   287   zdiv_zero ~> div_0
   288   zmod_zero ~> mod_0
   289   zdiv_zmod_equality ~> div_mod_equality2
   290   zdiv_zmod_equality2 ~> div_mod_equality
   291   zmod_zdiv_trivial ~> mod_div_trivial
   292   zdiv_zminus_zminus ~> div_minus_minus
   293   zmod_zminus_zminus ~> mod_minus_minus
   294   zdiv_zminus2 ~> div_minus_right
   295   zmod_zminus2 ~> mod_minus_right
   296   zdiv_minus1_right ~> div_minus1_right
   297   zmod_minus1_right ~> mod_minus1_right
   298   zdvd_mult_div_cancel ~> dvd_mult_div_cancel
   299   zmod_zmult1_eq ~> mod_mult_right_eq
   300   zpower_zmod ~> power_mod
   301   zdvd_zmod ~> dvd_mod
   302   zdvd_zmod_imp_zdvd ~> dvd_mod_imp_dvd
   303   mod_mult_distrib ~> mult_mod_left
   304   mod_mult_distrib2 ~> mult_mod_right
   305 
   306 * Removed redundant theorems nat_mult_2 and nat_mult_2_right; use
   307 generic mult_2 and mult_2_right instead. INCOMPATIBILITY.
   308 
   309 * Finite_Set.fold now qualified.  INCOMPATIBILITY.
   310 
   311 * Consolidated theorem names concerning fold combinators:
   312 
   313   inf_INFI_fold_inf ~> inf_INF_fold_inf
   314   sup_SUPR_fold_sup ~> sup_SUP_fold_sup
   315   INFI_fold_inf ~> INF_fold_inf
   316   SUPR_fold_sup ~> SUP_fold_sup
   317   union_set ~> union_set_fold
   318   minus_set ~> minus_set_fold
   319   INFI_set_fold ~> INF_set_fold
   320   SUPR_set_fold ~> SUP_set_fold
   321   INF_code ~> INF_set_foldr
   322   SUP_code ~> SUP_set_foldr
   323   foldr.simps ~> foldr.simps (in point-free formulation)
   324   foldr_fold_rev ~> foldr_conv_fold
   325   foldl_fold ~> foldl_conv_fold
   326   foldr_foldr ~> foldr_conv_foldl
   327   foldl_foldr ~> foldl_conv_foldr
   328   fold_set_remdups ~> fold_set_fold_remdups
   329   fold_set ~> fold_set_fold
   330   fold1_set ~> fold1_set_fold
   331 
   332 INCOMPATIBILITY.
   333 
   334 * Dropped rarely useful theorems concerning fold combinators:
   335 foldl_apply, foldl_fun_comm, foldl_rev, fold_weak_invariant,
   336 rev_foldl_cons, fold_set_remdups, fold_set, fold_set1,
   337 concat_conv_foldl, foldl_weak_invariant, foldl_invariant,
   338 foldr_invariant, foldl_absorb0, foldl_foldr1_lemma, foldl_foldr1,
   339 listsum_conv_fold, listsum_foldl, sort_foldl_insort, foldl_assoc,
   340 foldr_conv_foldl, start_le_sum, elem_le_sum, sum_eq_0_conv.
   341 INCOMPATIBILITY.  For the common phrases "%xs. List.foldr plus xs 0"
   342 and "List.foldl plus 0", prefer "List.listsum".  Otherwise it can be
   343 useful to boil down "List.foldr" and "List.foldl" to "List.fold" by
   344 unfolding "foldr_conv_fold" and "foldl_conv_fold".
   345 
   346 * Dropped lemmas minus_set_foldr, union_set_foldr, union_coset_foldr,
   347 inter_coset_foldr, Inf_fin_set_foldr, Sup_fin_set_foldr,
   348 Min_fin_set_foldr, Max_fin_set_foldr, Inf_set_foldr, Sup_set_foldr,
   349 INF_set_foldr, SUP_set_foldr.  INCOMPATIBILITY.  Prefer corresponding
   350 lemmas over fold rather than foldr, or make use of lemmas
   351 fold_conv_foldr and fold_rev.
   352 
   353 * Congruence rules Option.map_cong and Option.bind_cong for recursion
   354 through option types.
   355 
   356 * Concrete syntax for case expressions includes constraints for source
   357 positions, and thus produces Prover IDE markup for its bindings.
   358 INCOMPATIBILITY for old-style syntax translations that augment the
   359 pattern notation; e.g. see src/HOL/HOLCF/One.thy for translations of
   360 one_case.
   361 
   362 * Discontinued configuration option "syntax_positions": atomic terms
   363 in parse trees are always annotated by position constraints.
   364 
   365 * New theory HOL/Library/DAList provides an abstract type for
   366 association lists with distinct keys.
   367 
   368 * 'datatype' specifications allow explicit sort constraints.
   369 
   370 * Theory HOL/Library/Diagonalize has been removed. INCOMPATIBILITY,
   371 use theory HOL/Library/Nat_Bijection instead.
   372 
   373 * Theory HOL/Library/RBT_Impl: Backing implementation of red-black
   374 trees is now inside a type class context.  Names of affected
   375 operations and lemmas have been prefixed by rbt_.  INCOMPATIBILITY for
   376 theories working directly with raw red-black trees, adapt the names as
   377 follows:
   378 
   379   Operations:
   380   bulkload -> rbt_bulkload
   381   del_from_left -> rbt_del_from_left
   382   del_from_right -> rbt_del_from_right
   383   del -> rbt_del
   384   delete -> rbt_delete
   385   ins -> rbt_ins
   386   insert -> rbt_insert
   387   insertw -> rbt_insert_with
   388   insert_with_key -> rbt_insert_with_key
   389   map_entry -> rbt_map_entry
   390   lookup -> rbt_lookup
   391   sorted -> rbt_sorted
   392   tree_greater -> rbt_greater
   393   tree_less -> rbt_less
   394   tree_less_symbol -> rbt_less_symbol
   395   union -> rbt_union
   396   union_with -> rbt_union_with
   397   union_with_key -> rbt_union_with_key
   398 
   399   Lemmas:
   400   balance_left_sorted -> balance_left_rbt_sorted
   401   balance_left_tree_greater -> balance_left_rbt_greater
   402   balance_left_tree_less -> balance_left_rbt_less
   403   balance_right_sorted -> balance_right_rbt_sorted
   404   balance_right_tree_greater -> balance_right_rbt_greater
   405   balance_right_tree_less -> balance_right_rbt_less
   406   balance_sorted -> balance_rbt_sorted
   407   balance_tree_greater -> balance_rbt_greater
   408   balance_tree_less -> balance_rbt_less
   409   bulkload_is_rbt -> rbt_bulkload_is_rbt
   410   combine_sorted -> combine_rbt_sorted
   411   combine_tree_greater -> combine_rbt_greater
   412   combine_tree_less -> combine_rbt_less
   413   delete_in_tree -> rbt_delete_in_tree
   414   delete_is_rbt -> rbt_delete_is_rbt
   415   del_from_left_tree_greater -> rbt_del_from_left_rbt_greater
   416   del_from_left_tree_less -> rbt_del_from_left_rbt_less
   417   del_from_right_tree_greater -> rbt_del_from_right_rbt_greater
   418   del_from_right_tree_less -> rbt_del_from_right_rbt_less
   419   del_in_tree -> rbt_del_in_tree
   420   del_inv1_inv2 -> rbt_del_inv1_inv2
   421   del_sorted -> rbt_del_rbt_sorted
   422   del_tree_greater -> rbt_del_rbt_greater
   423   del_tree_less -> rbt_del_rbt_less
   424   dom_lookup_Branch -> dom_rbt_lookup_Branch
   425   entries_lookup -> entries_rbt_lookup
   426   finite_dom_lookup -> finite_dom_rbt_lookup
   427   insert_sorted -> rbt_insert_rbt_sorted
   428   insertw_is_rbt -> rbt_insertw_is_rbt
   429   insertwk_is_rbt -> rbt_insertwk_is_rbt
   430   insertwk_sorted -> rbt_insertwk_rbt_sorted
   431   insertw_sorted -> rbt_insertw_rbt_sorted
   432   ins_sorted -> ins_rbt_sorted
   433   ins_tree_greater -> ins_rbt_greater
   434   ins_tree_less -> ins_rbt_less
   435   is_rbt_sorted -> is_rbt_rbt_sorted
   436   lookup_balance -> rbt_lookup_balance
   437   lookup_bulkload -> rbt_lookup_rbt_bulkload
   438   lookup_delete -> rbt_lookup_rbt_delete
   439   lookup_Empty -> rbt_lookup_Empty
   440   lookup_from_in_tree -> rbt_lookup_from_in_tree
   441   lookup_in_tree -> rbt_lookup_in_tree
   442   lookup_ins -> rbt_lookup_ins
   443   lookup_insert -> rbt_lookup_rbt_insert
   444   lookup_insertw -> rbt_lookup_rbt_insertw
   445   lookup_insertwk -> rbt_lookup_rbt_insertwk
   446   lookup_keys -> rbt_lookup_keys
   447   lookup_map -> rbt_lookup_map
   448   lookup_map_entry -> rbt_lookup_rbt_map_entry
   449   lookup_tree_greater -> rbt_lookup_rbt_greater
   450   lookup_tree_less -> rbt_lookup_rbt_less
   451   lookup_union -> rbt_lookup_rbt_union
   452   map_entry_color_of -> rbt_map_entry_color_of
   453   map_entry_inv1 -> rbt_map_entry_inv1
   454   map_entry_inv2 -> rbt_map_entry_inv2
   455   map_entry_is_rbt -> rbt_map_entry_is_rbt
   456   map_entry_sorted -> rbt_map_entry_rbt_sorted
   457   map_entry_tree_greater -> rbt_map_entry_rbt_greater
   458   map_entry_tree_less -> rbt_map_entry_rbt_less
   459   map_tree_greater -> map_rbt_greater
   460   map_tree_less -> map_rbt_less
   461   map_sorted -> map_rbt_sorted
   462   paint_sorted -> paint_rbt_sorted
   463   paint_lookup -> paint_rbt_lookup
   464   paint_tree_greater -> paint_rbt_greater
   465   paint_tree_less -> paint_rbt_less
   466   sorted_entries -> rbt_sorted_entries
   467   tree_greater_eq_trans -> rbt_greater_eq_trans
   468   tree_greater_nit -> rbt_greater_nit
   469   tree_greater_prop -> rbt_greater_prop
   470   tree_greater_simps -> rbt_greater_simps
   471   tree_greater_trans -> rbt_greater_trans
   472   tree_less_eq_trans -> rbt_less_eq_trans
   473   tree_less_nit -> rbt_less_nit
   474   tree_less_prop -> rbt_less_prop
   475   tree_less_simps -> rbt_less_simps
   476   tree_less_trans -> rbt_less_trans
   477   tree_ord_props -> rbt_ord_props
   478   union_Branch -> rbt_union_Branch
   479   union_is_rbt -> rbt_union_is_rbt
   480   unionw_is_rbt -> rbt_unionw_is_rbt
   481   unionwk_is_rbt -> rbt_unionwk_is_rbt
   482   unionwk_sorted -> rbt_unionwk_rbt_sorted
   483 
   484 * Theory HOL/Library/Float: Floating point numbers are now defined as a
   485 subset of the real numbers. All operations are defined using the
   486 lifing-framework and proofs use the transfer method.
   487 INCOMPATIBILITY.
   488 
   489   Changed Operations:
   490   float_abs -> abs
   491   float_nprt -> nprt
   492   float_pprt -> pprt
   493   pow2 -> use powr
   494   round_down -> float_round_down
   495   round_up -> float_round_up
   496   scale -> exponent
   497 
   498   Removed Operations:
   499   ceiling_fl, lb_mult, lb_mod, ub_mult, ub_mod
   500 
   501   Renamed Lemmas:
   502   abs_float_def -> Float.compute_float_abs
   503   bitlen_ge0 -> bitlen_nonneg
   504   bitlen.simps -> Float.compute_bitlen
   505   float_components -> Float_mantissa_exponent
   506   float_divl.simps -> Float.compute_float_divl
   507   float_divr.simps -> Float.compute_float_divr
   508   float_eq_odd -> mult_powr_eq_mult_powr_iff
   509   float_power -> real_of_float_power
   510   lapprox_posrat_def -> Float.compute_lapprox_posrat
   511   lapprox_rat.simps -> Float.compute_lapprox_rat
   512   le_float_def' -> Float.compute_float_le
   513   le_float_def -> less_eq_float.rep_eq
   514   less_float_def' -> Float.compute_float_less
   515   less_float_def -> less_float.rep_eq
   516   normfloat_def -> Float.compute_normfloat
   517   normfloat_imp_odd_or_zero -> mantissa_not_dvd and mantissa_noteq_0
   518   normfloat -> normfloat_def
   519   normfloat_unique -> use normfloat_def
   520   number_of_float_Float -> Float.compute_float_numeral, Float.compute_float_neg_numeral
   521   one_float_def -> Float.compute_float_one
   522   plus_float_def -> Float.compute_float_plus
   523   rapprox_posrat_def -> Float.compute_rapprox_posrat
   524   rapprox_rat.simps -> Float.compute_rapprox_rat
   525   real_of_float_0 -> zero_float.rep_eq
   526   real_of_float_1 -> one_float.rep_eq
   527   real_of_float_abs -> abs_float.rep_eq
   528   real_of_float_add -> plus_float.rep_eq
   529   real_of_float_minus -> uminus_float.rep_eq
   530   real_of_float_mult -> times_float.rep_eq
   531   real_of_float_simp -> Float.rep_eq
   532   real_of_float_sub -> minus_float.rep_eq
   533   round_down.simps -> Float.compute_float_round_down
   534   round_up.simps -> Float.compute_float_round_up
   535   times_float_def -> Float.compute_float_times
   536   uminus_float_def -> Float.compute_float_uminus
   537   zero_float_def -> Float.compute_float_zero
   538 
   539   Lemmas not necessary anymore, use the transfer method:
   540   bitlen_B0, bitlen_B1, bitlen_ge1, bitlen_Min, bitlen_Pls, float_divl,
   541   float_divr, float_le_simp, float_less1_mantissa_bound,
   542   float_less_simp, float_less_zero, float_le_zero,
   543   float_pos_less1_e_neg, float_pos_m_pos, float_split, float_split2,
   544   floor_pos_exp, lapprox_posrat, lapprox_posrat_bottom, lapprox_rat,
   545   lapprox_rat_bottom, normalized_float, rapprox_posrat,
   546   rapprox_posrat_le1, rapprox_rat, real_of_float_ge0_exp,
   547   real_of_float_neg_exp, real_of_float_nge0_exp, round_down floor_fl,
   548   round_up, zero_le_float, zero_less_float
   549 
   550 * Session HOL-Word: Discontinued many redundant theorems specific to
   551 type 'a word. INCOMPATIBILITY, use the corresponding generic theorems
   552 instead.
   553 
   554   word_sub_alt ~> word_sub_wi
   555   word_add_alt ~> word_add_def
   556   word_mult_alt ~> word_mult_def
   557   word_minus_alt ~> word_minus_def
   558   word_0_alt ~> word_0_wi
   559   word_1_alt ~> word_1_wi
   560   word_add_0 ~> add_0_left
   561   word_add_0_right ~> add_0_right
   562   word_mult_1 ~> mult_1_left
   563   word_mult_1_right ~> mult_1_right
   564   word_add_commute ~> add_commute
   565   word_add_assoc ~> add_assoc
   566   word_add_left_commute ~> add_left_commute
   567   word_mult_commute ~> mult_commute
   568   word_mult_assoc ~> mult_assoc
   569   word_mult_left_commute ~> mult_left_commute
   570   word_left_distrib ~> left_distrib
   571   word_right_distrib ~> right_distrib
   572   word_left_minus ~> left_minus
   573   word_diff_0_right ~> diff_0_right
   574   word_diff_self ~> diff_self
   575   word_sub_def ~> diff_minus
   576   word_diff_minus ~> diff_minus
   577   word_add_ac ~> add_ac
   578   word_mult_ac ~> mult_ac
   579   word_plus_ac0 ~> add_0_left add_0_right add_ac
   580   word_times_ac1 ~> mult_1_left mult_1_right mult_ac
   581   word_order_trans ~> order_trans
   582   word_order_refl ~> order_refl
   583   word_order_antisym ~> order_antisym
   584   word_order_linear ~> linorder_linear
   585   lenw1_zero_neq_one ~> zero_neq_one
   586   word_number_of_eq ~> number_of_eq
   587   word_of_int_add_hom ~> wi_hom_add
   588   word_of_int_sub_hom ~> wi_hom_sub
   589   word_of_int_mult_hom ~> wi_hom_mult
   590   word_of_int_minus_hom ~> wi_hom_neg
   591   word_of_int_succ_hom ~> wi_hom_succ
   592   word_of_int_pred_hom ~> wi_hom_pred
   593   word_of_int_0_hom ~> word_0_wi
   594   word_of_int_1_hom ~> word_1_wi
   595 
   596 * New tactic "word_bitwise" for splitting machine word equalities and
   597 inequalities into logical circuits. Requires theory "WordBitwise" from HOL-Word
   598 session. Supports addition, subtraction, multiplication, shifting by
   599 constants, bitwise operators and numeric constants. Requires fixed-length word
   600 types, cannot operate on 'a word. Solves many standard word identies outright
   601 and converts more into first order problems amenable to blast or similar. See
   602 HOL/Word/WordBitwise.thy and examples in HOL/Word/Examples/WordExamples.thy.
   603 
   604 * Clarified attribute "mono_set": pure declaration without modifying
   605 the result of the fact expression.
   606 
   607 * "Transitive_Closure.ntrancl": bounded transitive closure on
   608 relations.
   609 
   610 * Constant "Set.not_member" now qualified.  INCOMPATIBILITY.
   611 
   612 * Theory Int: Discontinued many legacy theorems specific to type int.
   613 INCOMPATIBILITY, use the corresponding generic theorems instead.
   614 
   615   zminus_zminus ~> minus_minus
   616   zminus_0 ~> minus_zero
   617   zminus_zadd_distrib ~> minus_add_distrib
   618   zadd_commute ~> add_commute
   619   zadd_assoc ~> add_assoc
   620   zadd_left_commute ~> add_left_commute
   621   zadd_ac ~> add_ac
   622   zmult_ac ~> mult_ac
   623   zadd_0 ~> add_0_left
   624   zadd_0_right ~> add_0_right
   625   zadd_zminus_inverse2 ~> left_minus
   626   zmult_zminus ~> mult_minus_left
   627   zmult_commute ~> mult_commute
   628   zmult_assoc ~> mult_assoc
   629   zadd_zmult_distrib ~> left_distrib
   630   zadd_zmult_distrib2 ~> right_distrib
   631   zdiff_zmult_distrib ~> left_diff_distrib
   632   zdiff_zmult_distrib2 ~> right_diff_distrib
   633   zmult_1 ~> mult_1_left
   634   zmult_1_right ~> mult_1_right
   635   zle_refl ~> order_refl
   636   zle_trans ~> order_trans
   637   zle_antisym ~> order_antisym
   638   zle_linear ~> linorder_linear
   639   zless_linear ~> linorder_less_linear
   640   zadd_left_mono ~> add_left_mono
   641   zadd_strict_right_mono ~> add_strict_right_mono
   642   zadd_zless_mono ~> add_less_le_mono
   643   int_0_less_1 ~> zero_less_one
   644   int_0_neq_1 ~> zero_neq_one
   645   zless_le ~> less_le
   646   zpower_zadd_distrib ~> power_add
   647   zero_less_zpower_abs_iff ~> zero_less_power_abs_iff
   648   zero_le_zpower_abs ~> zero_le_power_abs
   649 
   650 * Theory Deriv: Renamed
   651 
   652   DERIV_nonneg_imp_nonincreasing ~> DERIV_nonneg_imp_nondecreasing
   653 
   654 * Theory Library/Multiset: Improved code generation of multisets.
   655 
   656 * Session HOL-Probability: Introduced the type "'a measure" to represent
   657 measures, this replaces the records 'a algebra and 'a measure_space. The
   658 locales based on subset_class now have two locale-parameters the space
   659 \<Omega> and the set of measurables sets M. The product of probability spaces
   660 uses now the same constant as the finite product of sigma-finite measure
   661 spaces "PiM :: ('i => 'a) measure". Most constants are defined now
   662 outside of locales and gain an additional parameter, like null_sets,
   663 almost_eventually or \<mu>'. Measure space constructions for distributions
   664 and densities now got their own constants distr and density. Instead of
   665 using locales to describe measure spaces with a finite space, the
   666 measure count_space and point_measure is introduced. INCOMPATIBILITY.
   667 
   668   Renamed constants:
   669   measure -> emeasure
   670   finite_measure.\<mu>' -> measure
   671   product_algebra_generator -> prod_algebra
   672   product_prob_space.emb -> prod_emb
   673   product_prob_space.infprod_algebra -> PiM
   674 
   675   Removed locales:
   676   completeable_measure_space
   677   finite_measure_space
   678   finite_prob_space
   679   finite_product_finite_prob_space
   680   finite_product_sigma_algebra
   681   finite_sigma_algebra
   682   measure_space
   683   pair_finite_prob_space
   684   pair_finite_sigma_algebra
   685   pair_finite_space
   686   pair_sigma_algebra
   687   product_sigma_algebra
   688 
   689   Removed constants:
   690   distribution -> use distr measure, or distributed predicate
   691   joint_distribution -> use distr measure, or distributed predicate
   692   product_prob_space.infprod_algebra -> use PiM
   693   subvimage
   694   image_space
   695   conditional_space
   696   pair_measure_generator
   697 
   698   Replacement theorems:
   699   sigma_algebra.measurable_sigma -> measurable_measure_of
   700   measure_space.additive -> emeasure_additive
   701   measure_space.measure_additive -> plus_emeasure
   702   measure_space.measure_mono -> emeasure_mono
   703   measure_space.measure_top -> emeasure_space
   704   measure_space.measure_compl -> emeasure_compl
   705   measure_space.measure_Diff -> emeasure_Diff
   706   measure_space.measure_countable_increasing -> emeasure_countable_increasing
   707   measure_space.continuity_from_below -> SUP_emeasure_incseq
   708   measure_space.measure_incseq -> incseq_emeasure
   709   measure_space.continuity_from_below_Lim -> Lim_emeasure_incseq
   710   measure_space.measure_decseq -> decseq_emeasure
   711   measure_space.continuity_from_above -> INF_emeasure_decseq
   712   measure_space.measure_insert -> emeasure_insert
   713   measure_space.measure_setsum -> setsum_emeasure
   714   measure_space.measure_finite_singleton -> emeasure_eq_setsum_singleton
   715   finite_additivity_sufficient -> ring_of_sets.countably_additiveI_finite
   716   measure_space.measure_setsum_split -> setsum_emeasure_cover
   717   measure_space.measure_subadditive -> subadditive
   718   measure_space.measure_subadditive_finite -> emeasure_subadditive_finite
   719   measure_space.measure_eq_0 -> emeasure_eq_0
   720   measure_space.measure_finitely_subadditive -> emeasure_subadditive_finite
   721   measure_space.measure_countably_subadditive -> emeasure_subadditive_countably
   722   measure_space.measure_UN_eq_0 -> emeasure_UN_eq_0
   723   measure_unique_Int_stable -> measure_eqI_generator_eq
   724   measure_space.measure_Diff_null_set -> emeasure_Diff_null_set
   725   measure_space.measure_Un_null_set -> emeasure_Un_null_set
   726   measure_space.almost_everywhere_def -> eventually_ae_filter
   727   measure_space.almost_everywhere_vimage -> AE_distrD
   728   measure_space.measure_space_vimage -> emeasure_distr
   729   measure_space.AE_iff_null_set -> AE_iff_null
   730   measure_space.real_measure_Union -> measure_Union
   731   measure_space.real_measure_finite_Union -> measure_finite_Union
   732   measure_space.real_measure_Diff -> measure_Diff
   733   measure_space.real_measure_UNION -> measure_UNION
   734   measure_space.real_measure_subadditive -> measure_subadditive
   735   measure_space.real_measure_setsum_singleton -> measure_eq_setsum_singleton
   736   measure_space.real_continuity_from_below -> Lim_measure_incseq
   737   measure_space.continuity_from_above_Lim -> Lim_emeasure_decseq
   738   measure_space.real_continuity_from_above -> Lim_measure_decseq
   739   measure_space.real_measure_countably_subadditive -> measure_subadditive_countably
   740   finite_measure.finite_measure -> finite_measure.emeasure_finite
   741   finite_measure.finite_measure_eq -> finite_measure.emeasure_eq_measure
   742   finite_measure.positive_measure' -> measure_nonneg
   743   finite_measure.real_measure -> finite_measure.emeasure_real
   744   finite_measure.empty_measure -> measure_empty
   745   finite_measure.finite_measure_countably_subadditive -> finite_measure.finite_measure_subadditive_countably
   746   finite_measure.finite_measure_finite_singleton -> finite_measure.finite_measure_eq_setsum_singleton
   747   finite_measure.finite_continuity_from_below -> finite_measure.finite_Lim_measure_incseq
   748   finite_measure.finite_continuity_from_above -> finite_measure.finite_Lim_measure_decseq
   749   measure_space.simple_integral_vimage -> simple_integral_distr
   750   measure_space.integrable_vimage -> integrable_distr
   751   measure_space.positive_integral_translated_density -> positive_integral_density
   752   measure_space.integral_translated_density -> integral_density
   753   measure_space.integral_vimage -> integral_distr
   754   measure_space_density -> emeasure_density
   755   measure_space.positive_integral_vimage -> positive_integral_distr
   756   measure_space.simple_function_vimage -> simple_function_comp
   757   measure_space.simple_integral_vimage -> simple_integral_distr
   758   pair_sigma_algebra.measurable_cut_fst -> sets_Pair1
   759   pair_sigma_algebra.measurable_cut_snd -> sets_Pair2
   760   pair_sigma_algebra.measurable_pair_image_fst -> measurable_Pair1
   761   pair_sigma_algebra.measurable_pair_image_snd -> measurable_Pair2
   762   pair_sigma_algebra.measurable_product_swap -> measurable_pair_swap_iff
   763   pair_sigma_finite.measure_cut_measurable_fst -> pair_sigma_finite.measurable_emeasure_Pair1
   764   pair_sigma_finite.measure_cut_measurable_snd -> pair_sigma_finite.measurable_emeasure_Pair2
   765   measure_space.measure_not_negative -> emeasure_not_MInf
   766   pair_sigma_finite.measure_preserving_swap -> pair_sigma_finite.distr_pair_swap
   767   pair_sigma_finite.pair_measure_alt -> pair_sigma_finite.emeasure_pair_measure_alt
   768   pair_sigma_finite.pair_measure_alt2 -> pair_sigma_finite.emeasure_pair_measure_alt2
   769   pair_sigma_finite.pair_measure_times -> pair_sigma_finite.emeasure_pair_measure_Times
   770   pair_sigma_algebra.pair_sigma_algebra_measurable -> measurable_pair_swap
   771   pair_sigma_algebra.pair_sigma_algebra_swap_measurable -> measurable_pair_swap'
   772   pair_sigma_algebra.sets_swap -> sets_pair_swap
   773   finite_product_sigma_algebra.in_P -> sets_PiM_I_finite
   774   Int_stable_product_algebra_generator -> positive_integral
   775   product_sigma_finite.measure_fold -> product_sigma_finite.distr_merge
   776   product_sigma_finite.measure_preserving_component_singelton -> product_sigma_finite.distr_singleton
   777   product_sigma_finite.measure_preserving_merge -> product_sigma_finite.distr_merge
   778   finite_product_sigma_algebra.P_empty -> space_PiM_empty, sets_PiM_empty
   779   product_algebra_generator_der -> prod_algebra_eq_finite
   780   product_algebra_generator_into_space -> prod_algebra_sets_into_space
   781   product_sigma_algebra.product_algebra_into_space -> space_closed
   782   product_algebraE -> prod_algebraE_all
   783   product_algebraI -> sets_PiM_I_finite
   784   product_measure_exists -> product_sigma_finite.sigma_finite
   785   sets_product_algebra -> sets_PiM
   786   sigma_product_algebra_sigma_eq -> sigma_prod_algebra_sigma_eq
   787   space_product_algebra -> space_PiM
   788   Int_stable_cuboids -> Int_stable_atLeastAtMost
   789   measure_space.density_is_absolutely_continuous -> absolutely_continuousI_density
   790   sigma_finite_measure.RN_deriv_vimage -> sigma_finite_measure.RN_deriv_distr
   791   prob_space_unique_Int_stable -> measure_eqI_prob_space
   792   sigma_finite_measure.disjoint_sigma_finite -> sigma_finite_disjoint
   793   prob_space.measure_space_1 -> prob_space.emeasure_space_1
   794   prob_space.prob_space_vimage -> prob_space_distr
   795   prob_space.random_variable_restrict -> measurable_restrict
   796   measure_preserving -> equality "distr M N f = N" "f : measurable M N"
   797   measure_unique_Int_stable_vimage -> measure_eqI_generator_eq
   798   measure_space.measure_preserving_Int_stable -> measure_eqI_generator_eq
   799   product_prob_space.finite_index_eq_finite_product -> product_prob_space.sets_PiM_generator
   800   product_prob_space.finite_measure_infprod_emb_Pi -> product_prob_space.measure_PiM_emb
   801   finite_product_prob_space.finite_measure_times -> finite_product_prob_space.finite_measure_PiM_emb
   802   product_prob_space.infprod_spec -> product_prob_space.emeasure_PiM_emb_not_empty
   803   product_prob_space.measurable_component -> measurable_component_singleton
   804   product_prob_space.measurable_emb -> measurable_prod_emb
   805   product_prob_space.measurable_into_infprod_algebra -> measurable_PiM_single
   806   product_prob_space.measurable_singleton_infprod -> measurable_component_singleton
   807   product_prob_space.measure_emb -> emeasure_prod_emb
   808   sequence_space.measure_infprod -> sequence_space.measure_PiM_countable
   809   product_prob_space.measure_preserving_restrict -> product_prob_space.distr_restrict
   810   prob_space.indep_distribution_eq_measure -> prob_space.indep_vars_iff_distr_eq_PiM
   811   prob_space.indep_var_distributionD -> prob_space.indep_var_distribution_eq
   812   conditional_entropy_positive -> conditional_entropy_nonneg_simple
   813   conditional_entropy_eq -> conditional_entropy_simple_distributed
   814   conditional_mutual_information_eq_mutual_information -> conditional_mutual_information_eq_mutual_information_simple
   815   conditional_mutual_information_generic_positive -> conditional_mutual_information_nonneg_simple
   816   conditional_mutual_information_positive -> conditional_mutual_information_nonneg_simple
   817   entropy_commute -> entropy_commute_simple
   818   entropy_eq -> entropy_simple_distributed
   819   entropy_generic_eq -> entropy_simple_distributed
   820   entropy_positive -> entropy_nonneg_simple
   821   entropy_uniform_max -> entropy_uniform
   822   KL_eq_0 -> KL_same_eq_0
   823   KL_eq_0_imp -> KL_eq_0_iff_eq
   824   KL_ge_0 -> KL_nonneg
   825   mutual_information_eq -> mutual_information_simple_distributed
   826   mutual_information_positive -> mutual_information_nonneg_simple
   827 
   828 * New "case_product" attribute to generate a case rule doing multiple
   829 case distinctions at the same time.  E.g.
   830 
   831   list.exhaust [case_product nat.exhaust]
   832 
   833 produces a rule which can be used to perform case distinction on both
   834 a list and a nat.
   835 
   836 * New Transfer package:
   837 
   838   - transfer_rule attribute: Maintains a collection of transfer rules,
   839     which relate constants at two different types. Transfer rules may
   840     relate different type instances of the same polymorphic constant,
   841     or they may relate an operation on a raw type to a corresponding
   842     operation on an abstract type (quotient or subtype). For example:
   843 
   844     ((A ===> B) ===> list_all2 A ===> list_all2 B) map map
   845     (cr_int ===> cr_int ===> cr_int) (%(x,y) (u,v). (x+u, y+v)) plus_int
   846 
   847   - transfer method: Replaces a subgoal on abstract types with an
   848     equivalent subgoal on the corresponding raw types. Constants are
   849     replaced with corresponding ones according to the transfer rules.
   850     Goals are generalized over all free variables by default; this is
   851     necessary for variables whose types changes, but can be overridden
   852     for specific variables with e.g. 'transfer fixing: x y z'.
   853     The variant transfer' method allows replacing a subgoal with
   854     one that is logically stronger (rather than equivalent).
   855 
   856   - relator_eq attribute: Collects identity laws for relators of
   857     various type constructors, e.g. "list_all2 (op =) = (op =)". The
   858     transfer method uses these lemmas to infer transfer rules for
   859     non-polymorphic constants on the fly.
   860 
   861   - transfer_prover method: Assists with proving a transfer rule for a
   862     new constant, provided the constant is defined in terms of other
   863     constants that already have transfer rules. It should be applied
   864     after unfolding the constant definitions.
   865 
   866   - HOL/ex/Transfer_Int_Nat.thy: Example theory demonstrating transfer
   867     from type nat to type int.
   868 
   869 * New Lifting package:
   870 
   871   - lift_definition command: Defines operations on an abstract type in
   872     terms of a corresponding operation on a representation type. Example
   873     syntax:
   874 
   875     lift_definition dlist_insert :: "'a => 'a dlist => 'a dlist"
   876       is List.insert
   877 
   878     Users must discharge a respectfulness proof obligation when each
   879     constant is defined. (For a type copy, i.e. a typedef with UNIV,
   880     the proof is discharged automatically.) The obligation is
   881     presented in a user-friendly, readable form; a respectfulness
   882     theorem in the standard format and a transfer rule are generated
   883     by the package.
   884 
   885   - Integration with code_abstype: For typedefs (e.g. subtypes
   886     corresponding to a datatype invariant, such as dlist),
   887     lift_definition generates a code certificate theorem and sets up
   888     code generation for each constant.
   889 
   890   - setup_lifting command: Sets up the Lifting package to work with
   891     a user-defined type. The user must provide either a quotient
   892     theorem or a type_definition theorem. The package configures
   893     transfer rules for equality and quantifiers on the type, and sets
   894     up the lift_definition command to work with the type.
   895 
   896   - Usage examples: See Quotient_Examples/Lift_DList.thy,
   897     Quotient_Examples/Lift_RBT.thy, Word/Word.thy and
   898     Library/Float.thy.
   899 
   900 * Quotient package:
   901 
   902   - The 'quotient_type' command now supports a 'morphisms' option with
   903     rep and abs functions, similar to typedef.
   904 
   905   - 'quotient_type' sets up new types to work with the Lifting and
   906     Transfer packages, as with 'setup_lifting'.
   907 
   908   - The 'quotient_definition' command now requires the user to prove a
   909     respectfulness property at the point where the constant is
   910     defined, similar to lift_definition; INCOMPATIBILITY.
   911 
   912   - Renamed predicate 'Quotient' to 'Quotient3', and renamed theorems
   913     accordingly, INCOMPATIBILITY.
   914 
   915 * New diagnostic command 'find_unused_assms' to find potentially
   916 superfluous assumptions in theorems using Quickcheck.
   917 
   918 * Quickcheck:
   919 
   920   - Quickcheck returns variable assignments as counterexamples, which
   921     allows to reveal the underspecification of functions under test.
   922     For example, refuting "hd xs = x", it presents the variable
   923     assignment xs = [] and x = a1 as a counterexample, assuming that
   924     any property is false whenever "hd []" occurs in it.
   925 
   926     These counterexample are marked as potentially spurious, as
   927     Quickcheck also returns "xs = []" as a counterexample to the
   928     obvious theorem "hd xs = hd xs".
   929 
   930     After finding a potentially spurious counterexample, Quickcheck
   931     continues searching for genuine ones.
   932 
   933     By default, Quickcheck shows potentially spurious and genuine
   934     counterexamples. The option "genuine_only" sets quickcheck to only
   935     show genuine counterexamples.
   936 
   937   - The command 'quickcheck_generator' creates random and exhaustive
   938     value generators for a given type and operations.
   939 
   940     It generates values by using the operations as if they were
   941     constructors of that type.
   942 
   943   - Support for multisets.
   944 
   945   - Added "use_subtype" options.
   946 
   947   - Added "quickcheck_locale" configuration to specify how to process
   948     conjectures in a locale context.
   949 
   950 * Nitpick:
   951   - Fixed infinite loop caused by the 'peephole_optim' option and
   952     affecting 'rat' and 'real'.
   953 
   954 * Sledgehammer:
   955   - Integrated more tightly with SPASS, as described in the ITP 2012 paper "More
   956     SPASS with Isabelle".
   957   - Made it try "smt" as a fallback if "metis" fails or times out.
   958   - Added support for the following provers: Alt-Ergo (via Why3 and TFF1),
   959     iProver, iProver-Eq.
   960   - Replaced remote E-SInE with remote Satallax in the default setup.
   961   - Sped up the minimizer.
   962   - Added "lam_trans", "uncurry_aliases", and "minimize" options.
   963   - Renamed "slicing" ("no_slicing") option to "slice" ("dont_slice").
   964   - Renamed "sound" option to "strict".
   965 
   966 * Metis:
   967   - Added possibility to specify lambda translations scheme as a
   968     parenthesized argument (e.g., "by (metis (lifting) ...)").
   969 
   970 * SMT:
   971   - Renamed "smt_fixed" option to "smt_read_only_certificates".
   972 
   973 * Command 'try0':
   974   - Renamed from 'try_methods'. INCOMPATIBILITY.
   975 
   976 * New "eventually_elim" method as a generalized variant of the
   977 eventually_elim* rules. Supports structured proofs.
   978 
   979 * HOL/TPTP: support to parse and import TPTP problems (all languages)
   980 into Isabelle/HOL.
   981 
   982 
   983 *** FOL ***
   984 
   985 * New "case_product" attribute (see HOL).
   986 
   987 
   988 *** ZF ***
   989 
   990 * Greater support for structured proofs involving induction or case
   991 analysis.
   992 
   993 * Much greater use of mathematical symbols.
   994 
   995 * Removal of many ML theorem bindings.  INCOMPATIBILITY.
   996 
   997 
   998 *** ML ***
   999 
  1000 * Antiquotation @{keyword "name"} produces a parser for outer syntax
  1001 from a minor keyword introduced via theory header declaration.
  1002 
  1003 * Antiquotation @{command_spec "name"} produces the
  1004 Outer_Syntax.command_spec from a major keyword introduced via theory
  1005 header declaration; it can be passed to Outer_Syntax.command etc.
  1006 
  1007 * Local_Theory.define no longer hard-wires default theorem name
  1008 "foo_def", but retains the binding as given.  If that is Binding.empty
  1009 / Attrib.empty_binding, the result is not registered as user-level
  1010 fact.  The Local_Theory.define_internal variant allows to specify a
  1011 non-empty name (used for the foundation in the background theory),
  1012 while omitting the fact binding in the user-context.  Potential
  1013 INCOMPATIBILITY for derived definitional packages: need to specify
  1014 naming policy for primitive definitions more explicitly.
  1015 
  1016 * Renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in
  1017 conformance with similar operations in structure Term and Logic.
  1018 
  1019 * Antiquotation @{attributes [...]} embeds attribute source
  1020 representation into the ML text, which is particularly useful with
  1021 declarations like Local_Theory.note.
  1022 
  1023 * Structure Proof_Context follows standard naming scheme.  Old
  1024 ProofContext has been discontinued.  INCOMPATIBILITY.
  1025 
  1026 * Refined Local_Theory.declaration {syntax, pervasive}, with subtle
  1027 change of semantics: update is applied to auxiliary local theory
  1028 context as well.
  1029 
  1030 * Modernized some old-style infix operations:
  1031 
  1032   addeqcongs    ~> Simplifier.add_eqcong
  1033   deleqcongs    ~> Simplifier.del_eqcong
  1034   addcongs      ~> Simplifier.add_cong
  1035   delcongs      ~> Simplifier.del_cong
  1036   setmksimps    ~> Simplifier.set_mksimps
  1037   setmkcong     ~> Simplifier.set_mkcong
  1038   setmksym      ~> Simplifier.set_mksym
  1039   setmkeqTrue   ~> Simplifier.set_mkeqTrue
  1040   settermless   ~> Simplifier.set_termless
  1041   setsubgoaler  ~> Simplifier.set_subgoaler
  1042   addsplits     ~> Splitter.add_split
  1043   delsplits     ~> Splitter.del_split
  1044 
  1045 
  1046 *** System ***
  1047 
  1048 * USER_HOME settings variable points to cross-platform user home
  1049 directory, which coincides with HOME on POSIX systems only.  Likewise,
  1050 the Isabelle path specification "~" now expands to $USER_HOME, instead
  1051 of former $HOME.  A different default for USER_HOME may be set
  1052 explicitly in shell environment, before Isabelle settings are
  1053 evaluated.  Minor INCOMPATIBILITY: need to adapt Isabelle path where
  1054 the generic user home was intended.
  1055 
  1056 * ISABELLE_JDK_HOME settings variable points to JDK with javac and jar
  1057 (not just JRE).
  1058 
  1059 * ISABELLE_HOME_WINDOWS refers to ISABELLE_HOME in windows file name
  1060 notation, which is useful for the jEdit file browser, for example.
  1061 
  1062 
  1063 
  1064 New in Isabelle2011-1 (October 2011)
  1065 ------------------------------------
  1066 
  1067 *** General ***
  1068 
  1069 * Improved Isabelle/jEdit Prover IDE (PIDE), which can be invoked as
  1070 "isabelle jedit" or "ISABELLE_HOME/Isabelle" on the command line.
  1071 
  1072   - Management of multiple theory files directly from the editor
  1073     buffer store -- bypassing the file-system (no requirement to save
  1074     files for checking).
  1075 
  1076   - Markup of formal entities within the text buffer, with semantic
  1077     highlighting, tooltips and hyperlinks to jump to defining source
  1078     positions.
  1079 
  1080   - Improved text rendering, with sub/superscripts in the source
  1081     buffer (including support for copy/paste wrt. output panel, HTML
  1082     theory output and other non-Isabelle text boxes).
  1083 
  1084   - Refined scheduling of proof checking and printing of results,
  1085     based on interactive editor view.  (Note: jEdit folding and
  1086     narrowing allows to restrict buffer perspectives explicitly.)
  1087 
  1088   - Reduced CPU performance requirements, usable on machines with few
  1089     cores.
  1090 
  1091   - Reduced memory requirements due to pruning of unused document
  1092     versions (garbage collection).
  1093 
  1094 See also ~~/src/Tools/jEdit/README.html for further information,
  1095 including some remaining limitations.
  1096 
  1097 * Theory loader: source files are exclusively located via the master
  1098 directory of each theory node (where the .thy file itself resides).
  1099 The global load path (such as src/HOL/Library) has been discontinued.
  1100 Note that the path element ~~ may be used to reference theories in the
  1101 Isabelle home folder -- for instance, "~~/src/HOL/Library/FuncSet".
  1102 INCOMPATIBILITY.
  1103 
  1104 * Theory loader: source files are identified by content via SHA1
  1105 digests.  Discontinued former path/modtime identification and optional
  1106 ISABELLE_FILE_IDENT plugin scripts.
  1107 
  1108 * Parallelization of nested Isar proofs is subject to
  1109 Goal.parallel_proofs_threshold (default 100).  See also isabelle
  1110 usedir option -Q.
  1111 
  1112 * Name space: former unsynchronized references are now proper
  1113 configuration options, with more conventional names:
  1114 
  1115   long_names   ~> names_long
  1116   short_names  ~> names_short
  1117   unique_names ~> names_unique
  1118 
  1119 Minor INCOMPATIBILITY, need to declare options in context like this:
  1120 
  1121   declare [[names_unique = false]]
  1122 
  1123 * Literal facts `prop` may contain dummy patterns, e.g. `_ = _`.  Note
  1124 that the result needs to be unique, which means fact specifications
  1125 may have to be refined after enriching a proof context.
  1126 
  1127 * Attribute "case_names" has been refined: the assumptions in each case
  1128 can be named now by following the case name with [name1 name2 ...].
  1129 
  1130 * Isabelle/Isar reference manual has been updated and extended:
  1131   - "Synopsis" provides a catalog of main Isar language concepts.
  1132   - Formal references in syntax diagrams, via @{rail} antiquotation.
  1133   - Updated material from classic "ref" manual, notably about
  1134     "Classical Reasoner".
  1135 
  1136 
  1137 *** HOL ***
  1138 
  1139 * Class bot and top require underlying partial order rather than
  1140 preorder: uniqueness of bot and top is guaranteed.  INCOMPATIBILITY.
  1141 
  1142 * Class complete_lattice: generalized a couple of lemmas from sets;
  1143 generalized theorems INF_cong and SUP_cong.  New type classes for
  1144 complete boolean algebras and complete linear orders.  Lemmas
  1145 Inf_less_iff, less_Sup_iff, INF_less_iff, less_SUP_iff now reside in
  1146 class complete_linorder.
  1147 
  1148 Changed proposition of lemmas Inf_bool_def, Sup_bool_def, Inf_fun_def,
  1149 Sup_fun_def, Inf_apply, Sup_apply.
  1150 
  1151 Removed redundant lemmas (the right hand side gives hints how to
  1152 replace them for (metis ...), or (simp only: ...) proofs):
  1153 
  1154   Inf_singleton ~> Inf_insert [where A="{}", unfolded Inf_empty inf_top_right]
  1155   Sup_singleton ~> Sup_insert [where A="{}", unfolded Sup_empty sup_bot_right]
  1156   Inf_binary ~> Inf_insert, Inf_empty, and inf_top_right
  1157   Sup_binary ~> Sup_insert, Sup_empty, and sup_bot_right
  1158   Int_eq_Inter ~> Inf_insert, Inf_empty, and inf_top_right
  1159   Un_eq_Union ~> Sup_insert, Sup_empty, and sup_bot_right
  1160   Inter_def ~> INF_def, image_def
  1161   Union_def ~> SUP_def, image_def
  1162   INT_eq ~> INF_def, and image_def
  1163   UN_eq ~> SUP_def, and image_def
  1164   INF_subset ~> INF_superset_mono [OF _ order_refl]
  1165 
  1166 More consistent and comprehensive names:
  1167 
  1168   INTER_eq_Inter_image ~> INF_def
  1169   UNION_eq_Union_image ~> SUP_def
  1170   INFI_def ~> INF_def
  1171   SUPR_def ~> SUP_def
  1172   INF_leI ~> INF_lower
  1173   INF_leI2 ~> INF_lower2
  1174   le_INFI ~> INF_greatest
  1175   le_SUPI ~> SUP_upper
  1176   le_SUPI2 ~> SUP_upper2
  1177   SUP_leI ~> SUP_least
  1178   INFI_bool_eq ~> INF_bool_eq
  1179   SUPR_bool_eq ~> SUP_bool_eq
  1180   INFI_apply ~> INF_apply
  1181   SUPR_apply ~> SUP_apply
  1182   INTER_def ~> INTER_eq
  1183   UNION_def ~> UNION_eq
  1184 
  1185 INCOMPATIBILITY.
  1186 
  1187 * Renamed theory Complete_Lattice to Complete_Lattices.
  1188 INCOMPATIBILITY.
  1189 
  1190 * Theory Complete_Lattices: lemmas Inf_eq_top_iff, INF_eq_top_iff,
  1191 INF_image, Inf_insert, INF_top, Inf_top_conv, INF_top_conv, SUP_bot,
  1192 Sup_bot_conv, SUP_bot_conv, Sup_eq_top_iff, SUP_eq_top_iff, SUP_image,
  1193 Sup_insert are now declared as [simp].  INCOMPATIBILITY.
  1194 
  1195 * Theory Lattice: lemmas compl_inf_bot, compl_le_comp_iff,
  1196 compl_sup_top, inf_idem, inf_left_idem, inf_sup_absorb, sup_idem,
  1197 sup_inf_absob, sup_left_idem are now declared as [simp].  Minor
  1198 INCOMPATIBILITY.
  1199 
  1200 * Added syntactic classes "inf" and "sup" for the respective
  1201 constants.  INCOMPATIBILITY: Changes in the argument order of the
  1202 (mostly internal) locale predicates for some derived classes.
  1203 
  1204 * Theorem collections ball_simps and bex_simps do not contain theorems
  1205 referring to UNION any longer; these have been moved to collection
  1206 UN_ball_bex_simps.  INCOMPATIBILITY.
  1207 
  1208 * Theory Archimedean_Field: floor now is defined as parameter of a
  1209 separate type class floor_ceiling.
  1210 
  1211 * Theory Finite_Set: more coherent development of fold_set locales:
  1212 
  1213     locale fun_left_comm ~> locale comp_fun_commute
  1214     locale fun_left_comm_idem ~> locale comp_fun_idem
  1215 
  1216 Both use point-free characterization; interpretation proofs may need
  1217 adjustment.  INCOMPATIBILITY.
  1218 
  1219 * Theory Limits: Type "'a net" has been renamed to "'a filter", in
  1220 accordance with standard mathematical terminology. INCOMPATIBILITY.
  1221 
  1222 * Theory Complex_Main: The locale interpretations for the
  1223 bounded_linear and bounded_bilinear locales have been removed, in
  1224 order to reduce the number of duplicate lemmas. Users must use the
  1225 original names for distributivity theorems, potential INCOMPATIBILITY.
  1226 
  1227   divide.add ~> add_divide_distrib
  1228   divide.diff ~> diff_divide_distrib
  1229   divide.setsum ~> setsum_divide_distrib
  1230   mult.add_right ~> right_distrib
  1231   mult.diff_right ~> right_diff_distrib
  1232   mult_right.setsum ~> setsum_right_distrib
  1233   mult_left.diff ~> left_diff_distrib
  1234 
  1235 * Theory Complex_Main: Several redundant theorems have been removed or
  1236 replaced by more general versions. INCOMPATIBILITY.
  1237 
  1238   real_diff_def ~> minus_real_def
  1239   real_divide_def ~> divide_real_def
  1240   real_less_def ~> less_le
  1241   real_abs_def ~> abs_real_def
  1242   real_sgn_def ~> sgn_real_def
  1243   real_mult_commute ~> mult_commute
  1244   real_mult_assoc ~> mult_assoc
  1245   real_mult_1 ~> mult_1_left
  1246   real_add_mult_distrib ~> left_distrib
  1247   real_zero_not_eq_one ~> zero_neq_one
  1248   real_mult_inverse_left ~> left_inverse
  1249   INVERSE_ZERO ~> inverse_zero
  1250   real_le_refl ~> order_refl
  1251   real_le_antisym ~> order_antisym
  1252   real_le_trans ~> order_trans
  1253   real_le_linear ~> linear
  1254   real_le_eq_diff ~> le_iff_diff_le_0
  1255   real_add_left_mono ~> add_left_mono
  1256   real_mult_order ~> mult_pos_pos
  1257   real_mult_less_mono2 ~> mult_strict_left_mono
  1258   real_of_int_real_of_nat ~> real_of_int_of_nat_eq
  1259   real_0_le_divide_iff ~> zero_le_divide_iff
  1260   realpow_two_disj ~> power2_eq_iff
  1261   real_squared_diff_one_factored ~> square_diff_one_factored
  1262   realpow_two_diff ~> square_diff_square_factored
  1263   reals_complete2 ~> complete_real
  1264   real_sum_squared_expand ~> power2_sum
  1265   exp_ln_eq ~> ln_unique
  1266   expi_add ~> exp_add
  1267   expi_zero ~> exp_zero
  1268   lemma_DERIV_subst ~> DERIV_cong
  1269   LIMSEQ_Zfun_iff ~> tendsto_Zfun_iff
  1270   LIMSEQ_const ~> tendsto_const
  1271   LIMSEQ_norm ~> tendsto_norm
  1272   LIMSEQ_add ~> tendsto_add
  1273   LIMSEQ_minus ~> tendsto_minus
  1274   LIMSEQ_minus_cancel ~> tendsto_minus_cancel
  1275   LIMSEQ_diff ~> tendsto_diff
  1276   bounded_linear.LIMSEQ ~> bounded_linear.tendsto
  1277   bounded_bilinear.LIMSEQ ~> bounded_bilinear.tendsto
  1278   LIMSEQ_mult ~> tendsto_mult
  1279   LIMSEQ_inverse ~> tendsto_inverse
  1280   LIMSEQ_divide ~> tendsto_divide
  1281   LIMSEQ_pow ~> tendsto_power
  1282   LIMSEQ_setsum ~> tendsto_setsum
  1283   LIMSEQ_setprod ~> tendsto_setprod
  1284   LIMSEQ_norm_zero ~> tendsto_norm_zero_iff
  1285   LIMSEQ_rabs_zero ~> tendsto_rabs_zero_iff
  1286   LIMSEQ_imp_rabs ~> tendsto_rabs
  1287   LIMSEQ_add_minus ~> tendsto_add [OF _ tendsto_minus]
  1288   LIMSEQ_add_const ~> tendsto_add [OF _ tendsto_const]
  1289   LIMSEQ_diff_const ~> tendsto_diff [OF _ tendsto_const]
  1290   LIMSEQ_Complex ~> tendsto_Complex
  1291   LIM_ident ~> tendsto_ident_at
  1292   LIM_const ~> tendsto_const
  1293   LIM_add ~> tendsto_add
  1294   LIM_add_zero ~> tendsto_add_zero
  1295   LIM_minus ~> tendsto_minus
  1296   LIM_diff ~> tendsto_diff
  1297   LIM_norm ~> tendsto_norm
  1298   LIM_norm_zero ~> tendsto_norm_zero
  1299   LIM_norm_zero_cancel ~> tendsto_norm_zero_cancel
  1300   LIM_norm_zero_iff ~> tendsto_norm_zero_iff
  1301   LIM_rabs ~> tendsto_rabs
  1302   LIM_rabs_zero ~> tendsto_rabs_zero
  1303   LIM_rabs_zero_cancel ~> tendsto_rabs_zero_cancel
  1304   LIM_rabs_zero_iff ~> tendsto_rabs_zero_iff
  1305   LIM_compose ~> tendsto_compose
  1306   LIM_mult ~> tendsto_mult
  1307   LIM_scaleR ~> tendsto_scaleR
  1308   LIM_of_real ~> tendsto_of_real
  1309   LIM_power ~> tendsto_power
  1310   LIM_inverse ~> tendsto_inverse
  1311   LIM_sgn ~> tendsto_sgn
  1312   isCont_LIM_compose ~> isCont_tendsto_compose
  1313   bounded_linear.LIM ~> bounded_linear.tendsto
  1314   bounded_linear.LIM_zero ~> bounded_linear.tendsto_zero
  1315   bounded_bilinear.LIM ~> bounded_bilinear.tendsto
  1316   bounded_bilinear.LIM_prod_zero ~> bounded_bilinear.tendsto_zero
  1317   bounded_bilinear.LIM_left_zero ~> bounded_bilinear.tendsto_left_zero
  1318   bounded_bilinear.LIM_right_zero ~> bounded_bilinear.tendsto_right_zero
  1319   LIM_inverse_fun ~> tendsto_inverse [OF tendsto_ident_at]
  1320 
  1321 * Theory Complex_Main: The definition of infinite series was
  1322 generalized.  Now it is defined on the type class {topological_space,
  1323 comm_monoid_add}.  Hence it is useable also for extended real numbers.
  1324 
  1325 * Theory Complex_Main: The complex exponential function "expi" is now
  1326 a type-constrained abbreviation for "exp :: complex => complex"; thus
  1327 several polymorphic lemmas about "exp" are now applicable to "expi".
  1328 
  1329 * Code generation:
  1330 
  1331   - Theory Library/Code_Char_ord provides native ordering of
  1332     characters in the target language.
  1333 
  1334   - Commands code_module and code_library are legacy, use export_code
  1335     instead.
  1336 
  1337   - Method "evaluation" is legacy, use method "eval" instead.
  1338 
  1339   - Legacy evaluator "SML" is deactivated by default.  May be
  1340     reactivated by the following theory command:
  1341 
  1342       setup {* Value.add_evaluator ("SML", Codegen.eval_term) *}
  1343 
  1344 * Declare ext [intro] by default.  Rare INCOMPATIBILITY.
  1345 
  1346 * New proof method "induction" that gives induction hypotheses the
  1347 name "IH", thus distinguishing them from further hypotheses that come
  1348 from rule induction.  The latter are still called "hyps".  Method
  1349 "induction" is a thin wrapper around "induct" and follows the same
  1350 syntax.
  1351 
  1352 * Method "fastsimp" has been renamed to "fastforce", but "fastsimp" is
  1353 still available as a legacy feature for some time.
  1354 
  1355 * Nitpick:
  1356   - Added "need" and "total_consts" options.
  1357   - Reintroduced "show_skolems" option by popular demand.
  1358   - Renamed attribute: nitpick_def ~> nitpick_unfold.
  1359     INCOMPATIBILITY.
  1360 
  1361 * Sledgehammer:
  1362   - Use quasi-sound (and efficient) translations by default.
  1363   - Added support for the following provers: E-ToFoF, LEO-II,
  1364     Satallax, SNARK, Waldmeister, and Z3 with TPTP syntax.
  1365   - Automatically preplay and minimize proofs before showing them if
  1366     this can be done within reasonable time.
  1367   - sledgehammer available_provers ~> sledgehammer supported_provers.
  1368     INCOMPATIBILITY.
  1369   - Added "preplay_timeout", "slicing", "type_enc", "sound",
  1370     "max_mono_iters", and "max_new_mono_instances" options.
  1371   - Removed "explicit_apply" and "full_types" options as well as "Full
  1372     Types" Proof General menu item. INCOMPATIBILITY.
  1373 
  1374 * Metis:
  1375   - Removed "metisF" -- use "metis" instead. INCOMPATIBILITY.
  1376   - Obsoleted "metisFT" -- use "metis (full_types)" instead.
  1377     INCOMPATIBILITY.
  1378 
  1379 * Command 'try':
  1380   - Renamed 'try_methods' and added "simp:", "intro:", "dest:", and
  1381     "elim:" options. INCOMPATIBILITY.
  1382   - Introduced 'try' that not only runs 'try_methods' but also
  1383     'solve_direct', 'sledgehammer', 'quickcheck', and 'nitpick'.
  1384 
  1385 * Quickcheck:
  1386   - Added "eval" option to evaluate terms for the found counterexample
  1387     (currently only supported by the default (exhaustive) tester).
  1388   - Added post-processing of terms to obtain readable counterexamples
  1389     (currently only supported by the default (exhaustive) tester).
  1390   - New counterexample generator quickcheck[narrowing] enables
  1391     narrowing-based testing.  Requires the Glasgow Haskell compiler
  1392     with its installation location defined in the Isabelle settings
  1393     environment as ISABELLE_GHC.
  1394   - Removed quickcheck tester "SML" based on the SML code generator
  1395     (formly in HOL/Library).
  1396 
  1397 * Function package: discontinued option "tailrec".  INCOMPATIBILITY,
  1398 use 'partial_function' instead.
  1399 
  1400 * Theory Library/Extended_Reals replaces now the positive extended
  1401 reals found in probability theory. This file is extended by
  1402 Multivariate_Analysis/Extended_Real_Limits.
  1403 
  1404 * Theory Library/Old_Recdef: old 'recdef' package has been moved here,
  1405 from where it must be imported explicitly if it is really required.
  1406 INCOMPATIBILITY.
  1407 
  1408 * Theory Library/Wfrec: well-founded recursion combinator "wfrec" has
  1409 been moved here.  INCOMPATIBILITY.
  1410 
  1411 * Theory Library/Saturated provides type of numbers with saturated
  1412 arithmetic.
  1413 
  1414 * Theory Library/Product_Lattice defines a pointwise ordering for the
  1415 product type 'a * 'b, and provides instance proofs for various order
  1416 and lattice type classes.
  1417 
  1418 * Theory Library/Countable now provides the "countable_datatype" proof
  1419 method for proving "countable" class instances for datatypes.
  1420 
  1421 * Theory Library/Cset_Monad allows do notation for computable sets
  1422 (cset) via the generic monad ad-hoc overloading facility.
  1423 
  1424 * Library: Theories of common data structures are split into theories
  1425 for implementation, an invariant-ensuring type, and connection to an
  1426 abstract type. INCOMPATIBILITY.
  1427 
  1428   - RBT is split into RBT and RBT_Mapping.
  1429   - AssocList is split and renamed into AList and AList_Mapping.
  1430   - DList is split into DList_Impl, DList, and DList_Cset.
  1431   - Cset is split into Cset and List_Cset.
  1432 
  1433 * Theory Library/Nat_Infinity has been renamed to
  1434 Library/Extended_Nat, with name changes of the following types and
  1435 constants:
  1436 
  1437   type inat   ~> type enat
  1438   Fin         ~> enat
  1439   Infty       ~> infinity (overloaded)
  1440   iSuc        ~> eSuc
  1441   the_Fin     ~> the_enat
  1442 
  1443 Every theorem name containing "inat", "Fin", "Infty", or "iSuc" has
  1444 been renamed accordingly. INCOMPATIBILITY.
  1445 
  1446 * Session Multivariate_Analysis: The euclidean_space type class now
  1447 fixes a constant "Basis :: 'a set" consisting of the standard
  1448 orthonormal basis for the type. Users now have the option of
  1449 quantifying over this set instead of using the "basis" function, e.g.
  1450 "ALL x:Basis. P x" vs "ALL i<DIM('a). P (basis i)".
  1451 
  1452 * Session Multivariate_Analysis: Type "('a, 'b) cart" has been renamed
  1453 to "('a, 'b) vec" (the syntax "'a ^ 'b" remains unaffected). Constants
  1454 "Cart_nth" and "Cart_lambda" have been respectively renamed to
  1455 "vec_nth" and "vec_lambda"; theorems mentioning those names have
  1456 changed to match. Definition theorems for overloaded constants now use
  1457 the standard "foo_vec_def" naming scheme. A few other theorems have
  1458 been renamed as follows (INCOMPATIBILITY):
  1459 
  1460   Cart_eq          ~> vec_eq_iff
  1461   dist_nth_le_cart ~> dist_vec_nth_le
  1462   tendsto_vector   ~> vec_tendstoI
  1463   Cauchy_vector    ~> vec_CauchyI
  1464 
  1465 * Session Multivariate_Analysis: Several duplicate theorems have been
  1466 removed, and other theorems have been renamed or replaced with more
  1467 general versions. INCOMPATIBILITY.
  1468 
  1469   finite_choice ~> finite_set_choice
  1470   eventually_conjI ~> eventually_conj
  1471   eventually_and ~> eventually_conj_iff
  1472   eventually_false ~> eventually_False
  1473   setsum_norm ~> norm_setsum
  1474   Lim_sequentially ~> LIMSEQ_def
  1475   Lim_ident_at ~> LIM_ident
  1476   Lim_const ~> tendsto_const
  1477   Lim_cmul ~> tendsto_scaleR [OF tendsto_const]
  1478   Lim_neg ~> tendsto_minus
  1479   Lim_add ~> tendsto_add
  1480   Lim_sub ~> tendsto_diff
  1481   Lim_mul ~> tendsto_scaleR
  1482   Lim_vmul ~> tendsto_scaleR [OF _ tendsto_const]
  1483   Lim_null_norm ~> tendsto_norm_zero_iff [symmetric]
  1484   Lim_linear ~> bounded_linear.tendsto
  1485   Lim_component ~> tendsto_euclidean_component
  1486   Lim_component_cart ~> tendsto_vec_nth
  1487   Lim_inner ~> tendsto_inner [OF tendsto_const]
  1488   dot_lsum ~> inner_setsum_left
  1489   dot_rsum ~> inner_setsum_right
  1490   continuous_cmul ~> continuous_scaleR [OF continuous_const]
  1491   continuous_neg ~> continuous_minus
  1492   continuous_sub ~> continuous_diff
  1493   continuous_vmul ~> continuous_scaleR [OF _ continuous_const]
  1494   continuous_mul ~> continuous_scaleR
  1495   continuous_inv ~> continuous_inverse
  1496   continuous_at_within_inv ~> continuous_at_within_inverse
  1497   continuous_at_inv ~> continuous_at_inverse
  1498   continuous_at_norm ~> continuous_norm [OF continuous_at_id]
  1499   continuous_at_infnorm ~> continuous_infnorm [OF continuous_at_id]
  1500   continuous_at_component ~> continuous_component [OF continuous_at_id]
  1501   continuous_on_neg ~> continuous_on_minus
  1502   continuous_on_sub ~> continuous_on_diff
  1503   continuous_on_cmul ~> continuous_on_scaleR [OF continuous_on_const]
  1504   continuous_on_vmul ~> continuous_on_scaleR [OF _ continuous_on_const]
  1505   continuous_on_mul ~> continuous_on_scaleR
  1506   continuous_on_mul_real ~> continuous_on_mult
  1507   continuous_on_inner ~> continuous_on_inner [OF continuous_on_const]
  1508   continuous_on_norm ~> continuous_on_norm [OF continuous_on_id]
  1509   continuous_on_inverse ~> continuous_on_inv
  1510   uniformly_continuous_on_neg ~> uniformly_continuous_on_minus
  1511   uniformly_continuous_on_sub ~> uniformly_continuous_on_diff
  1512   subset_interior ~> interior_mono
  1513   subset_closure ~> closure_mono
  1514   closure_univ ~> closure_UNIV
  1515   real_arch_lt ~> reals_Archimedean2
  1516   real_arch ~> reals_Archimedean3
  1517   real_abs_norm ~> abs_norm_cancel
  1518   real_abs_sub_norm ~> norm_triangle_ineq3
  1519   norm_cauchy_schwarz_abs ~> Cauchy_Schwarz_ineq2
  1520 
  1521 * Session HOL-Probability:
  1522   - Caratheodory's extension lemma is now proved for ring_of_sets.
  1523   - Infinite products of probability measures are now available.
  1524   - Sigma closure is independent, if the generator is independent
  1525   - Use extended reals instead of positive extended
  1526     reals. INCOMPATIBILITY.
  1527 
  1528 * Session HOLCF: Discontinued legacy theorem names, INCOMPATIBILITY.
  1529 
  1530   expand_fun_below ~> fun_below_iff
  1531   below_fun_ext ~> fun_belowI
  1532   expand_cfun_eq ~> cfun_eq_iff
  1533   ext_cfun ~> cfun_eqI
  1534   expand_cfun_below ~> cfun_below_iff
  1535   below_cfun_ext ~> cfun_belowI
  1536   monofun_fun_fun ~> fun_belowD
  1537   monofun_fun_arg ~> monofunE
  1538   monofun_lub_fun ~> adm_monofun [THEN admD]
  1539   cont_lub_fun ~> adm_cont [THEN admD]
  1540   cont2cont_Rep_CFun ~> cont2cont_APP
  1541   cont_Rep_CFun_app ~> cont_APP_app
  1542   cont_Rep_CFun_app_app ~> cont_APP_app_app
  1543   cont_cfun_fun ~> cont_Rep_cfun1 [THEN contE]
  1544   cont_cfun_arg ~> cont_Rep_cfun2 [THEN contE]
  1545   contlub_cfun ~> lub_APP [symmetric]
  1546   contlub_LAM ~> lub_LAM [symmetric]
  1547   thelubI ~> lub_eqI
  1548   UU_I ~> bottomI
  1549   lift_distinct1 ~> lift.distinct(1)
  1550   lift_distinct2 ~> lift.distinct(2)
  1551   Def_not_UU ~> lift.distinct(2)
  1552   Def_inject ~> lift.inject
  1553   below_UU_iff ~> below_bottom_iff
  1554   eq_UU_iff ~> eq_bottom_iff
  1555 
  1556 
  1557 *** Document preparation ***
  1558 
  1559 * Antiquotation @{rail} layouts railroad syntax diagrams, see also
  1560 isar-ref manual, both for description and actual application of the
  1561 same.
  1562 
  1563 * Antiquotation @{value} evaluates the given term and presents its
  1564 result.
  1565 
  1566 * Antiquotations: term style "isub" provides ad-hoc conversion of
  1567 variables x1, y23 into subscripted form x\<^isub>1,
  1568 y\<^isub>2\<^isub>3.
  1569 
  1570 * Predefined LaTeX macros for Isabelle symbols \<bind> and \<then>
  1571 (e.g. see ~~/src/HOL/Library/Monad_Syntax.thy).
  1572 
  1573 * Localized \isabellestyle switch can be used within blocks or groups
  1574 like this:
  1575 
  1576   \isabellestyle{it}  %preferred default
  1577   {\isabellestylett @{text "typewriter stuff"}}
  1578 
  1579 * Discontinued special treatment of hard tabulators.  Implicit
  1580 tab-width is now defined as 1.  Potential INCOMPATIBILITY for visual
  1581 layouts.
  1582 
  1583 
  1584 *** ML ***
  1585 
  1586 * The inner syntax of sort/type/term/prop supports inlined YXML
  1587 representations within quoted string tokens.  By encoding logical
  1588 entities via Term_XML (in ML or Scala) concrete syntax can be
  1589 bypassed, which is particularly useful for producing bits of text
  1590 under external program control.
  1591 
  1592 * Antiquotations for ML and document preparation are managed as theory
  1593 data, which requires explicit setup.
  1594 
  1595 * Isabelle_Process.is_active allows tools to check if the official
  1596 process wrapper is running (Isabelle/Scala/jEdit) or the old TTY loop
  1597 (better known as Proof General).
  1598 
  1599 * Structure Proof_Context follows standard naming scheme.  Old
  1600 ProofContext is still available for some time as legacy alias.
  1601 
  1602 * Structure Timing provides various operations for timing; supersedes
  1603 former start_timing/end_timing etc.
  1604 
  1605 * Path.print is the official way to show file-system paths to users
  1606 (including quotes etc.).
  1607 
  1608 * Inner syntax: identifiers in parse trees of generic categories
  1609 "logic", "aprop", "idt" etc. carry position information (disguised as
  1610 type constraints).  Occasional INCOMPATIBILITY with non-compliant
  1611 translations that choke on unexpected type constraints.  Positions can
  1612 be stripped in ML translations via Syntax.strip_positions /
  1613 Syntax.strip_positions_ast, or via the syntax constant
  1614 "_strip_positions" within parse trees.  As last resort, positions can
  1615 be disabled via the configuration option Syntax.positions, which is
  1616 called "syntax_positions" in Isar attribute syntax.
  1617 
  1618 * Discontinued special status of various ML structures that contribute
  1619 to structure Syntax (Ast, Lexicon, Mixfix, Parser, Printer etc.): less
  1620 pervasive content, no inclusion in structure Syntax.  INCOMPATIBILITY,
  1621 refer directly to Ast.Constant, Lexicon.is_identifier,
  1622 Syntax_Trans.mk_binder_tr etc.
  1623 
  1624 * Typed print translation: discontinued show_sorts argument, which is
  1625 already available via context of "advanced" translation.
  1626 
  1627 * Refined PARALLEL_GOALS tactical: degrades gracefully for schematic
  1628 goal states; body tactic needs to address all subgoals uniformly.
  1629 
  1630 * Slightly more special eq_list/eq_set, with shortcut involving
  1631 pointer equality (assumes that eq relation is reflexive).
  1632 
  1633 * Classical tactics use proper Proof.context instead of historic types
  1634 claset/clasimpset.  Old-style declarations like addIs, addEs, addDs
  1635 operate directly on Proof.context.  Raw type claset retains its use as
  1636 snapshot of the classical context, which can be recovered via
  1637 (put_claset HOL_cs) etc.  Type clasimpset has been discontinued.
  1638 INCOMPATIBILITY, classical tactics and derived proof methods require
  1639 proper Proof.context.
  1640 
  1641 
  1642 *** System ***
  1643 
  1644 * Discontinued support for Poly/ML 5.2, which was the last version
  1645 without proper multithreading and TimeLimit implementation.
  1646 
  1647 * Discontinued old lib/scripts/polyml-platform, which has been
  1648 obsolete since Isabelle2009-2.
  1649 
  1650 * Various optional external tools are referenced more robustly and
  1651 uniformly by explicit Isabelle settings as follows:
  1652 
  1653   ISABELLE_CSDP   (formerly CSDP_EXE)
  1654   ISABELLE_GHC    (formerly EXEC_GHC or GHC_PATH)
  1655   ISABELLE_OCAML  (formerly EXEC_OCAML)
  1656   ISABELLE_SWIPL  (formerly EXEC_SWIPL)
  1657   ISABELLE_YAP    (formerly EXEC_YAP)
  1658 
  1659 Note that automated detection from the file-system or search path has
  1660 been discontinued.  INCOMPATIBILITY.
  1661 
  1662 * Scala layer provides JVM method invocation service for static
  1663 methods of type (String)String, see Invoke_Scala.method in ML.  For
  1664 example:
  1665 
  1666   Invoke_Scala.method "java.lang.System.getProperty" "java.home"
  1667 
  1668 Together with YXML.string_of_body/parse_body and XML.Encode/Decode
  1669 this allows to pass structured values between ML and Scala.
  1670 
  1671 * The IsabelleText fonts includes some further glyphs to support the
  1672 Prover IDE.  Potential INCOMPATIBILITY: users who happen to have
  1673 installed a local copy (which is normally *not* required) need to
  1674 delete or update it from ~~/lib/fonts/.
  1675 
  1676 
  1677 
  1678 New in Isabelle2011 (January 2011)
  1679 ----------------------------------
  1680 
  1681 *** General ***
  1682 
  1683 * Experimental Prover IDE based on Isabelle/Scala and jEdit (see
  1684 src/Tools/jEdit).  This also serves as IDE for Isabelle/ML, with
  1685 useful tooltips and hyperlinks produced from its static analysis.  The
  1686 bundled component provides an executable Isabelle tool that can be run
  1687 like this:
  1688 
  1689   Isabelle2011/bin/isabelle jedit
  1690 
  1691 * Significantly improved Isabelle/Isar implementation manual.
  1692 
  1693 * System settings: ISABELLE_HOME_USER now includes ISABELLE_IDENTIFIER
  1694 (and thus refers to something like $HOME/.isabelle/Isabelle2011),
  1695 while the default heap location within that directory lacks that extra
  1696 suffix.  This isolates multiple Isabelle installations from each
  1697 other, avoiding problems with old settings in new versions.
  1698 INCOMPATIBILITY, need to copy/upgrade old user settings manually.
  1699 
  1700 * Source files are always encoded as UTF-8, instead of old-fashioned
  1701 ISO-Latin-1.  INCOMPATIBILITY.  Isabelle LaTeX documents might require
  1702 the following package declarations:
  1703 
  1704   \usepackage[utf8]{inputenc}
  1705   \usepackage{textcomp}
  1706 
  1707 * Explicit treatment of UTF-8 sequences as Isabelle symbols, such that
  1708 a Unicode character is treated as a single symbol, not a sequence of
  1709 non-ASCII bytes as before.  Since Isabelle/ML string literals may
  1710 contain symbols without further backslash escapes, Unicode can now be
  1711 used here as well.  Recall that Symbol.explode in ML provides a
  1712 consistent view on symbols, while raw explode (or String.explode)
  1713 merely give a byte-oriented representation.
  1714 
  1715 * Theory loader: source files are primarily located via the master
  1716 directory of each theory node (where the .thy file itself resides).
  1717 The global load path is still partially available as legacy feature.
  1718 Minor INCOMPATIBILITY due to subtle change in file lookup: use
  1719 explicit paths, relatively to the theory.
  1720 
  1721 * Special treatment of ML file names has been discontinued.
  1722 Historically, optional extensions .ML or .sml were added on demand --
  1723 at the cost of clarity of file dependencies.  Recall that Isabelle/ML
  1724 files exclusively use the .ML extension.  Minor INCOMPATIBILTY.
  1725 
  1726 * Various options that affect pretty printing etc. are now properly
  1727 handled within the context via configuration options, instead of
  1728 unsynchronized references or print modes.  There are both ML Config.T
  1729 entities and Isar declaration attributes to access these.
  1730 
  1731   ML (Config.T)                 Isar (attribute)
  1732 
  1733   eta_contract                  eta_contract
  1734   show_brackets                 show_brackets
  1735   show_sorts                    show_sorts
  1736   show_types                    show_types
  1737   show_question_marks           show_question_marks
  1738   show_consts                   show_consts
  1739   show_abbrevs                  show_abbrevs
  1740 
  1741   Syntax.ast_trace              syntax_ast_trace
  1742   Syntax.ast_stat               syntax_ast_stat
  1743   Syntax.ambiguity_level        syntax_ambiguity_level
  1744 
  1745   Goal_Display.goals_limit      goals_limit
  1746   Goal_Display.show_main_goal   show_main_goal
  1747 
  1748   Method.rule_trace             rule_trace
  1749 
  1750   Thy_Output.display            thy_output_display
  1751   Thy_Output.quotes             thy_output_quotes
  1752   Thy_Output.indent             thy_output_indent
  1753   Thy_Output.source             thy_output_source
  1754   Thy_Output.break              thy_output_break
  1755 
  1756 Note that corresponding "..._default" references in ML may only be
  1757 changed globally at the ROOT session setup, but *not* within a theory.
  1758 The option "show_abbrevs" supersedes the former print mode
  1759 "no_abbrevs" with inverted meaning.
  1760 
  1761 * More systematic naming of some configuration options.
  1762 INCOMPATIBILITY.
  1763 
  1764   trace_simp  ~>  simp_trace
  1765   debug_simp  ~>  simp_debug
  1766 
  1767 * Support for real valued configuration options, using simplistic
  1768 floating-point notation that coincides with the inner syntax for
  1769 float_token.
  1770 
  1771 * Support for real valued preferences (with approximative PGIP type):
  1772 front-ends need to accept "pgint" values in float notation.
  1773 INCOMPATIBILITY.
  1774 
  1775 * The IsabelleText font now includes Cyrillic, Hebrew, Arabic from
  1776 DejaVu Sans.
  1777 
  1778 * Discontinued support for Poly/ML 5.0 and 5.1 versions.
  1779 
  1780 
  1781 *** Pure ***
  1782 
  1783 * Command 'type_synonym' (with single argument) replaces somewhat
  1784 outdated 'types', which is still available as legacy feature for some
  1785 time.
  1786 
  1787 * Command 'nonterminal' (with 'and' separated list of arguments)
  1788 replaces somewhat outdated 'nonterminals'.  INCOMPATIBILITY.
  1789 
  1790 * Command 'notepad' replaces former 'example_proof' for
  1791 experimentation in Isar without any result.  INCOMPATIBILITY.
  1792 
  1793 * Locale interpretation commands 'interpret' and 'sublocale' accept
  1794 lists of equations to map definitions in a locale to appropriate
  1795 entities in the context of the interpretation.  The 'interpretation'
  1796 command already provided this functionality.
  1797 
  1798 * Diagnostic command 'print_dependencies' prints the locale instances
  1799 that would be activated if the specified expression was interpreted in
  1800 the current context.  Variant "print_dependencies!" assumes a context
  1801 without interpretations.
  1802 
  1803 * Diagnostic command 'print_interps' prints interpretations in proofs
  1804 in addition to interpretations in theories.
  1805 
  1806 * Discontinued obsolete 'global' and 'local' commands to manipulate
  1807 the theory name space.  Rare INCOMPATIBILITY.  The ML functions
  1808 Sign.root_path and Sign.local_path may be applied directly where this
  1809 feature is still required for historical reasons.
  1810 
  1811 * Discontinued obsolete 'constdefs' command.  INCOMPATIBILITY, use
  1812 'definition' instead.
  1813 
  1814 * The "prems" fact, which refers to the accidental collection of
  1815 foundational premises in the context, is now explicitly marked as
  1816 legacy feature and will be discontinued soon.  Consider using "assms"
  1817 of the head statement or reference facts by explicit names.
  1818 
  1819 * Document antiquotations @{class} and @{type} print classes and type
  1820 constructors.
  1821 
  1822 * Document antiquotation @{file} checks file/directory entries within
  1823 the local file system.
  1824 
  1825 
  1826 *** HOL ***
  1827 
  1828 * Coercive subtyping: functions can be declared as coercions and type
  1829 inference will add them as necessary upon input of a term.  Theory
  1830 Complex_Main declares real :: nat => real and real :: int => real as
  1831 coercions. A coercion function f is declared like this:
  1832 
  1833   declare [[coercion f]]
  1834 
  1835 To lift coercions through type constructors (e.g. from nat => real to
  1836 nat list => real list), map functions can be declared, e.g.
  1837 
  1838   declare [[coercion_map map]]
  1839 
  1840 Currently coercion inference is activated only in theories including
  1841 real numbers, i.e. descendants of Complex_Main.  This is controlled by
  1842 the configuration option "coercion_enabled", e.g. it can be enabled in
  1843 other theories like this:
  1844 
  1845   declare [[coercion_enabled]]
  1846 
  1847 * Command 'partial_function' provides basic support for recursive
  1848 function definitions over complete partial orders.  Concrete instances
  1849 are provided for i) the option type, ii) tail recursion on arbitrary
  1850 types, and iii) the heap monad of Imperative_HOL.  See
  1851 src/HOL/ex/Fundefs.thy and src/HOL/Imperative_HOL/ex/Linked_Lists.thy
  1852 for examples.
  1853 
  1854 * Function package: f.psimps rules are no longer implicitly declared
  1855 as [simp].  INCOMPATIBILITY.
  1856 
  1857 * Datatype package: theorems generated for executable equality (class
  1858 "eq") carry proper names and are treated as default code equations.
  1859 
  1860 * Inductive package: now offers command 'inductive_simps' to
  1861 automatically derive instantiated and simplified equations for
  1862 inductive predicates, similar to 'inductive_cases'.
  1863 
  1864 * Command 'enriched_type' allows to register properties of the
  1865 functorial structure of types.
  1866 
  1867 * Improved infrastructure for term evaluation using code generator
  1868 techniques, in particular static evaluation conversions.
  1869 
  1870 * Code generator: Scala (2.8 or higher) has been added to the target
  1871 languages.
  1872 
  1873 * Code generator: globbing constant expressions "*" and "Theory.*"
  1874 have been replaced by the more idiomatic "_" and "Theory._".
  1875 INCOMPATIBILITY.
  1876 
  1877 * Code generator: export_code without explicit file declaration prints
  1878 to standard output.  INCOMPATIBILITY.
  1879 
  1880 * Code generator: do not print function definitions for case
  1881 combinators any longer.
  1882 
  1883 * Code generator: simplification with rules determined with
  1884 src/Tools/Code/code_simp.ML and method "code_simp".
  1885 
  1886 * Code generator for records: more idiomatic representation of record
  1887 types.  Warning: records are not covered by ancient SML code
  1888 generation any longer.  INCOMPATIBILITY.  In cases of need, a suitable
  1889 rep_datatype declaration helps to succeed then:
  1890 
  1891   record 'a foo = ...
  1892   ...
  1893   rep_datatype foo_ext ...
  1894 
  1895 * Records: logical foundation type for records does not carry a
  1896 '_type' suffix any longer (obsolete due to authentic syntax).
  1897 INCOMPATIBILITY.
  1898 
  1899 * Quickcheck now by default uses exhaustive testing instead of random
  1900 testing.  Random testing can be invoked by "quickcheck [random]",
  1901 exhaustive testing by "quickcheck [exhaustive]".
  1902 
  1903 * Quickcheck instantiates polymorphic types with small finite
  1904 datatypes by default. This enables a simple execution mechanism to
  1905 handle quantifiers and function equality over the finite datatypes.
  1906 
  1907 * Quickcheck random generator has been renamed from "code" to
  1908 "random".  INCOMPATIBILITY.
  1909 
  1910 * Quickcheck now has a configurable time limit which is set to 30
  1911 seconds by default. This can be changed by adding [timeout = n] to the
  1912 quickcheck command. The time limit for Auto Quickcheck is still set
  1913 independently.
  1914 
  1915 * Quickcheck in locales considers interpretations of that locale for
  1916 counter example search.
  1917 
  1918 * Sledgehammer:
  1919   - Added "smt" and "remote_smt" provers based on the "smt" proof
  1920     method. See the Sledgehammer manual for details ("isabelle doc
  1921     sledgehammer").
  1922   - Renamed commands:
  1923     sledgehammer atp_info ~> sledgehammer running_provers
  1924     sledgehammer atp_kill ~> sledgehammer kill_provers
  1925     sledgehammer available_atps ~> sledgehammer available_provers
  1926     INCOMPATIBILITY.
  1927   - Renamed options:
  1928     sledgehammer [atps = ...] ~> sledgehammer [provers = ...]
  1929     sledgehammer [atp = ...] ~> sledgehammer [prover = ...]
  1930     sledgehammer [timeout = 77 s] ~> sledgehammer [timeout = 77]
  1931     (and "ms" and "min" are no longer supported)
  1932     INCOMPATIBILITY.
  1933 
  1934 * Nitpick:
  1935   - Renamed options:
  1936     nitpick [timeout = 77 s] ~> nitpick [timeout = 77]
  1937     nitpick [tac_timeout = 777 ms] ~> nitpick [tac_timeout = 0.777]
  1938     INCOMPATIBILITY.
  1939   - Added support for partial quotient types.
  1940   - Added local versions of the "Nitpick.register_xxx" functions.
  1941   - Added "whack" option.
  1942   - Allow registration of quotient types as codatatypes.
  1943   - Improved "merge_type_vars" option to merge more types.
  1944   - Removed unsound "fast_descrs" option.
  1945   - Added custom symmetry breaking for datatypes, making it possible to reach
  1946     higher cardinalities.
  1947   - Prevent the expansion of too large definitions.
  1948 
  1949 * Proof methods "metis" and "meson" now have configuration options
  1950 "meson_trace", "metis_trace", and "metis_verbose" that can be enabled
  1951 to diagnose these tools. E.g.
  1952 
  1953     using [[metis_trace = true]]
  1954 
  1955 * Auto Solve: Renamed "Auto Solve Direct".  The tool is now available
  1956 manually as command 'solve_direct'.
  1957 
  1958 * The default SMT solver Z3 must be enabled explicitly (due to
  1959 licensing issues) by setting the environment variable
  1960 Z3_NON_COMMERCIAL in etc/settings of the component, for example.  For
  1961 commercial applications, the SMT solver CVC3 is provided as fall-back;
  1962 changing the SMT solver is done via the configuration option
  1963 "smt_solver".
  1964 
  1965 * Remote SMT solvers need to be referred to by the "remote_" prefix,
  1966 i.e. "remote_cvc3" and "remote_z3".
  1967 
  1968 * Added basic SMT support for datatypes, records, and typedefs using
  1969 the oracle mode (no proofs).  Direct support of pairs has been dropped
  1970 in exchange (pass theorems fst_conv snd_conv pair_collapse to the SMT
  1971 support for a similar behavior).  Minor INCOMPATIBILITY.
  1972 
  1973 * Changed SMT configuration options:
  1974   - Renamed:
  1975     z3_proofs ~> smt_oracle (with inverted meaning)
  1976     z3_trace_assms ~> smt_trace_used_facts
  1977     INCOMPATIBILITY.
  1978   - Added:
  1979     smt_verbose
  1980     smt_random_seed
  1981     smt_datatypes
  1982     smt_infer_triggers
  1983     smt_monomorph_limit
  1984     cvc3_options
  1985     remote_cvc3_options
  1986     remote_z3_options
  1987     yices_options
  1988 
  1989 * Boogie output files (.b2i files) need to be declared in the theory
  1990 header.
  1991 
  1992 * Simplification procedure "list_to_set_comprehension" rewrites list
  1993 comprehensions applied to List.set to set comprehensions.  Occasional
  1994 INCOMPATIBILITY, may be deactivated like this:
  1995 
  1996   declare [[simproc del: list_to_set_comprehension]]
  1997 
  1998 * Removed old version of primrec package.  INCOMPATIBILITY.
  1999 
  2000 * Removed simplifier congruence rule of "prod_case", as has for long
  2001 been the case with "split".  INCOMPATIBILITY.
  2002 
  2003 * String.literal is a type, but not a datatype.  INCOMPATIBILITY.
  2004 
  2005 * Removed [split_format ... and ... and ...] version of
  2006 [split_format].  Potential INCOMPATIBILITY.
  2007 
  2008 * Predicate "sorted" now defined inductively, with nice induction
  2009 rules.  INCOMPATIBILITY: former sorted.simps now named sorted_simps.
  2010 
  2011 * Constant "contents" renamed to "the_elem", to free the generic name
  2012 contents for other uses.  INCOMPATIBILITY.
  2013 
  2014 * Renamed class eq and constant eq (for code generation) to class
  2015 equal and constant equal, plus renaming of related facts and various
  2016 tuning.  INCOMPATIBILITY.
  2017 
  2018 * Dropped type classes mult_mono and mult_mono1.  INCOMPATIBILITY.
  2019 
  2020 * Removed output syntax "'a ~=> 'b" for "'a => 'b option".
  2021 INCOMPATIBILITY.
  2022 
  2023 * Renamed theory Fset to Cset, type Fset.fset to Cset.set, in order to
  2024 avoid confusion with finite sets.  INCOMPATIBILITY.
  2025 
  2026 * Abandoned locales equiv, congruent and congruent2 for equivalence
  2027 relations.  INCOMPATIBILITY: use equivI rather than equiv_intro (same
  2028 for congruent(2)).
  2029 
  2030 * Some previously unqualified names have been qualified:
  2031 
  2032   types
  2033     bool ~> HOL.bool
  2034     nat ~> Nat.nat
  2035 
  2036   constants
  2037     Trueprop ~> HOL.Trueprop
  2038     True ~> HOL.True
  2039     False ~> HOL.False
  2040     op & ~> HOL.conj
  2041     op | ~> HOL.disj
  2042     op --> ~> HOL.implies
  2043     op = ~> HOL.eq
  2044     Not ~> HOL.Not
  2045     The ~> HOL.The
  2046     All ~> HOL.All
  2047     Ex ~> HOL.Ex
  2048     Ex1 ~> HOL.Ex1
  2049     Let ~> HOL.Let
  2050     If ~> HOL.If
  2051     Ball ~> Set.Ball
  2052     Bex ~> Set.Bex
  2053     Suc ~> Nat.Suc
  2054     Pair ~> Product_Type.Pair
  2055     fst ~> Product_Type.fst
  2056     snd ~> Product_Type.snd
  2057     curry ~> Product_Type.curry
  2058     op : ~> Set.member
  2059     Collect ~> Set.Collect
  2060 
  2061 INCOMPATIBILITY.
  2062 
  2063 * More canonical naming convention for some fundamental definitions:
  2064 
  2065     bot_bool_eq ~> bot_bool_def
  2066     top_bool_eq ~> top_bool_def
  2067     inf_bool_eq ~> inf_bool_def
  2068     sup_bool_eq ~> sup_bool_def
  2069     bot_fun_eq  ~> bot_fun_def
  2070     top_fun_eq  ~> top_fun_def
  2071     inf_fun_eq  ~> inf_fun_def
  2072     sup_fun_eq  ~> sup_fun_def
  2073 
  2074 INCOMPATIBILITY.
  2075 
  2076 * More stylized fact names:
  2077 
  2078   expand_fun_eq ~> fun_eq_iff
  2079   expand_set_eq ~> set_eq_iff
  2080   set_ext       ~> set_eqI
  2081   nat_number    ~> eval_nat_numeral
  2082 
  2083 INCOMPATIBILITY.
  2084 
  2085 * Refactoring of code-generation specific operations in theory List:
  2086 
  2087   constants
  2088     null ~> List.null
  2089 
  2090   facts
  2091     mem_iff ~> member_def
  2092     null_empty ~> null_def
  2093 
  2094 INCOMPATIBILITY.  Note that these were not supposed to be used
  2095 regularly unless for striking reasons; their main purpose was code
  2096 generation.
  2097 
  2098 Various operations from the Haskell prelude are used for generating
  2099 Haskell code.
  2100 
  2101 * Term "bij f" is now an abbreviation of "bij_betw f UNIV UNIV".  Term
  2102 "surj f" is now an abbreviation of "range f = UNIV".  The theorems
  2103 bij_def and surj_def are unchanged.  INCOMPATIBILITY.
  2104 
  2105 * Abolished some non-alphabetic type names: "prod" and "sum" replace
  2106 "*" and "+" respectively.  INCOMPATIBILITY.
  2107 
  2108 * Name "Plus" of disjoint sum operator "<+>" is now hidden.  Write
  2109 "Sum_Type.Plus" instead.
  2110 
  2111 * Constant "split" has been merged with constant "prod_case"; names of
  2112 ML functions, facts etc. involving split have been retained so far,
  2113 though.  INCOMPATIBILITY.
  2114 
  2115 * Dropped old infix syntax "_ mem _" for List.member; use "_ : set _"
  2116 instead.  INCOMPATIBILITY.
  2117 
  2118 * Removed lemma "Option.is_none_none" which duplicates "is_none_def".
  2119 INCOMPATIBILITY.
  2120 
  2121 * Former theory Library/Enum is now part of the HOL-Main image.
  2122 INCOMPATIBILITY: all constants of the Enum theory now have to be
  2123 referred to by its qualified name.
  2124 
  2125   enum    ~>  Enum.enum
  2126   nlists  ~>  Enum.nlists
  2127   product ~>  Enum.product
  2128 
  2129 * Theory Library/Monad_Syntax provides do-syntax for monad types.
  2130 Syntax in Library/State_Monad has been changed to avoid ambiguities.
  2131 INCOMPATIBILITY.
  2132 
  2133 * Theory Library/SetsAndFunctions has been split into
  2134 Library/Function_Algebras and Library/Set_Algebras; canonical names
  2135 for instance definitions for functions; various improvements.
  2136 INCOMPATIBILITY.
  2137 
  2138 * Theory Library/Multiset provides stable quicksort implementation of
  2139 sort_key.
  2140 
  2141 * Theory Library/Multiset: renamed empty_idemp ~> empty_neutral.
  2142 INCOMPATIBILITY.
  2143 
  2144 * Session Multivariate_Analysis: introduced a type class for euclidean
  2145 space.  Most theorems are now stated in terms of euclidean spaces
  2146 instead of finite cartesian products.
  2147 
  2148   types
  2149     real ^ 'n ~>  'a::real_vector
  2150               ~>  'a::euclidean_space
  2151               ~>  'a::ordered_euclidean_space
  2152         (depends on your needs)
  2153 
  2154   constants
  2155      _ $ _        ~> _ $$ _
  2156      \<chi> x. _  ~> \<chi>\<chi> x. _
  2157      CARD('n)     ~> DIM('a)
  2158 
  2159 Also note that the indices are now natural numbers and not from some
  2160 finite type. Finite cartesian products of euclidean spaces, products
  2161 of euclidean spaces the real and complex numbers are instantiated to
  2162 be euclidean_spaces.  INCOMPATIBILITY.
  2163 
  2164 * Session Probability: introduced pextreal as positive extended real
  2165 numbers.  Use pextreal as value for measures.  Introduce the
  2166 Radon-Nikodym derivative, product spaces and Fubini's theorem for
  2167 arbitrary sigma finite measures.  Introduces Lebesgue measure based on
  2168 the integral in Multivariate Analysis.  INCOMPATIBILITY.
  2169 
  2170 * Session Imperative_HOL: revamped, corrected dozens of inadequacies.
  2171 INCOMPATIBILITY.
  2172 
  2173 * Session SPARK (with image HOL-SPARK) provides commands to load and
  2174 prove verification conditions generated by the SPARK Ada program
  2175 verifier.  See also src/HOL/SPARK and src/HOL/SPARK/Examples.
  2176 
  2177 
  2178 *** HOL-Algebra ***
  2179 
  2180 * Theorems for additive ring operations (locale abelian_monoid and
  2181 descendants) are generated by interpretation from their multiplicative
  2182 counterparts.  Names (in particular theorem names) have the mandatory
  2183 qualifier 'add'.  Previous theorem names are redeclared for
  2184 compatibility.
  2185 
  2186 * Structure "int_ring" is now an abbreviation (previously a
  2187 definition).  This fits more natural with advanced interpretations.
  2188 
  2189 
  2190 *** HOLCF ***
  2191 
  2192 * The domain package now runs in definitional mode by default: The
  2193 former command 'new_domain' is now called 'domain'.  To use the domain
  2194 package in its original axiomatic mode, use 'domain (unsafe)'.
  2195 INCOMPATIBILITY.
  2196 
  2197 * The new class "domain" is now the default sort.  Class "predomain"
  2198 is an unpointed version of "domain". Theories can be updated by
  2199 replacing sort annotations as shown below.  INCOMPATIBILITY.
  2200 
  2201   'a::type ~> 'a::countable
  2202   'a::cpo  ~> 'a::predomain
  2203   'a::pcpo ~> 'a::domain
  2204 
  2205 * The old type class "rep" has been superseded by class "domain".
  2206 Accordingly, users of the definitional package must remove any
  2207 "default_sort rep" declarations.  INCOMPATIBILITY.
  2208 
  2209 * The domain package (definitional mode) now supports unpointed
  2210 predomain argument types, as long as they are marked 'lazy'. (Strict
  2211 arguments must be in class "domain".) For example, the following
  2212 domain definition now works:
  2213 
  2214   domain natlist = nil | cons (lazy "nat discr") (lazy "natlist")
  2215 
  2216 * Theory HOLCF/Library/HOL_Cpo provides cpo and predomain class
  2217 instances for types from main HOL: bool, nat, int, char, 'a + 'b,
  2218 'a option, and 'a list.  Additionally, it configures fixrec and the
  2219 domain package to work with these types.  For example:
  2220 
  2221   fixrec isInl :: "('a + 'b) u -> tr"
  2222     where "isInl$(up$(Inl x)) = TT" | "isInl$(up$(Inr y)) = FF"
  2223 
  2224   domain V = VFun (lazy "V -> V") | VCon (lazy "nat") (lazy "V list")
  2225 
  2226 * The "(permissive)" option of fixrec has been replaced with a
  2227 per-equation "(unchecked)" option. See
  2228 src/HOL/HOLCF/Tutorial/Fixrec_ex.thy for examples. INCOMPATIBILITY.
  2229 
  2230 * The "bifinite" class no longer fixes a constant "approx"; the class
  2231 now just asserts that such a function exists.  INCOMPATIBILITY.
  2232 
  2233 * Former type "alg_defl" has been renamed to "defl".  HOLCF no longer
  2234 defines an embedding of type 'a defl into udom by default; instances
  2235 of "bifinite" and "domain" classes are available in
  2236 src/HOL/HOLCF/Library/Defl_Bifinite.thy.
  2237 
  2238 * The syntax "REP('a)" has been replaced with "DEFL('a)".
  2239 
  2240 * The predicate "directed" has been removed.  INCOMPATIBILITY.
  2241 
  2242 * The type class "finite_po" has been removed.  INCOMPATIBILITY.
  2243 
  2244 * The function "cprod_map" has been renamed to "prod_map".
  2245 INCOMPATIBILITY.
  2246 
  2247 * The monadic bind operator on each powerdomain has new binder syntax
  2248 similar to sets, e.g. "\<Union>\<sharp>x\<in>xs. t" represents
  2249 "upper_bind\<cdot>xs\<cdot>(\<Lambda> x. t)".
  2250 
  2251 * The infix syntax for binary union on each powerdomain has changed
  2252 from e.g. "+\<sharp>" to "\<union>\<sharp>", for consistency with set
  2253 syntax.  INCOMPATIBILITY.
  2254 
  2255 * The constant "UU" has been renamed to "bottom".  The syntax "UU" is
  2256 still supported as an input translation.
  2257 
  2258 * Renamed some theorems (the original names are also still available).
  2259 
  2260   expand_fun_below   ~> fun_below_iff
  2261   below_fun_ext      ~> fun_belowI
  2262   expand_cfun_eq     ~> cfun_eq_iff
  2263   ext_cfun           ~> cfun_eqI
  2264   expand_cfun_below  ~> cfun_below_iff
  2265   below_cfun_ext     ~> cfun_belowI
  2266   cont2cont_Rep_CFun ~> cont2cont_APP
  2267 
  2268 * The Abs and Rep functions for various types have changed names.
  2269 Related theorem names have also changed to match. INCOMPATIBILITY.
  2270 
  2271   Rep_CFun  ~> Rep_cfun
  2272   Abs_CFun  ~> Abs_cfun
  2273   Rep_Sprod ~> Rep_sprod
  2274   Abs_Sprod ~> Abs_sprod
  2275   Rep_Ssum  ~> Rep_ssum
  2276   Abs_Ssum  ~> Abs_ssum
  2277 
  2278 * Lemmas with names of the form *_defined_iff or *_strict_iff have
  2279 been renamed to *_bottom_iff.  INCOMPATIBILITY.
  2280 
  2281 * Various changes to bisimulation/coinduction with domain package:
  2282 
  2283   - Definitions of "bisim" constants no longer mention definedness.
  2284   - With mutual recursion, "bisim" predicate is now curried.
  2285   - With mutual recursion, each type gets a separate coind theorem.
  2286   - Variable names in bisim_def and coinduct rules have changed.
  2287 
  2288 INCOMPATIBILITY.
  2289 
  2290 * Case combinators generated by the domain package for type "foo" are
  2291 now named "foo_case" instead of "foo_when".  INCOMPATIBILITY.
  2292 
  2293 * Several theorems have been renamed to more accurately reflect the
  2294 names of constants and types involved.  INCOMPATIBILITY.
  2295 
  2296   thelub_const    ~> lub_const
  2297   lub_const       ~> is_lub_const
  2298   thelubI         ~> lub_eqI
  2299   is_lub_lub      ~> is_lubD2
  2300   lubI            ~> is_lub_lub
  2301   unique_lub      ~> is_lub_unique
  2302   is_ub_lub       ~> is_lub_rangeD1
  2303   lub_bin_chain   ~> is_lub_bin_chain
  2304   lub_fun         ~> is_lub_fun
  2305   thelub_fun      ~> lub_fun
  2306   thelub_cfun     ~> lub_cfun
  2307   thelub_Pair     ~> lub_Pair
  2308   lub_cprod       ~> is_lub_prod
  2309   thelub_cprod    ~> lub_prod
  2310   minimal_cprod   ~> minimal_prod
  2311   inst_cprod_pcpo ~> inst_prod_pcpo
  2312   UU_I            ~> bottomI
  2313   compact_UU      ~> compact_bottom
  2314   deflation_UU    ~> deflation_bottom
  2315   finite_deflation_UU ~> finite_deflation_bottom
  2316 
  2317 * Many legacy theorem names have been discontinued.  INCOMPATIBILITY.
  2318 
  2319   sq_ord_less_eq_trans ~> below_eq_trans
  2320   sq_ord_eq_less_trans ~> eq_below_trans
  2321   refl_less            ~> below_refl
  2322   trans_less           ~> below_trans
  2323   antisym_less         ~> below_antisym
  2324   antisym_less_inverse ~> po_eq_conv [THEN iffD1]
  2325   box_less             ~> box_below
  2326   rev_trans_less       ~> rev_below_trans
  2327   not_less2not_eq      ~> not_below2not_eq
  2328   less_UU_iff          ~> below_UU_iff
  2329   flat_less_iff        ~> flat_below_iff
  2330   adm_less             ~> adm_below
  2331   adm_not_less         ~> adm_not_below
  2332   adm_compact_not_less ~> adm_compact_not_below
  2333   less_fun_def         ~> below_fun_def
  2334   expand_fun_less      ~> fun_below_iff
  2335   less_fun_ext         ~> fun_belowI
  2336   less_discr_def       ~> below_discr_def
  2337   discr_less_eq        ~> discr_below_eq
  2338   less_unit_def        ~> below_unit_def
  2339   less_cprod_def       ~> below_prod_def
  2340   prod_lessI           ~> prod_belowI
  2341   Pair_less_iff        ~> Pair_below_iff
  2342   fst_less_iff         ~> fst_below_iff
  2343   snd_less_iff         ~> snd_below_iff
  2344   expand_cfun_less     ~> cfun_below_iff
  2345   less_cfun_ext        ~> cfun_belowI
  2346   injection_less       ~> injection_below
  2347   less_up_def          ~> below_up_def
  2348   not_Iup_less         ~> not_Iup_below
  2349   Iup_less             ~> Iup_below
  2350   up_less              ~> up_below
  2351   Def_inject_less_eq   ~> Def_below_Def
  2352   Def_less_is_eq       ~> Def_below_iff
  2353   spair_less_iff       ~> spair_below_iff
  2354   less_sprod           ~> below_sprod
  2355   spair_less           ~> spair_below
  2356   sfst_less_iff        ~> sfst_below_iff
  2357   ssnd_less_iff        ~> ssnd_below_iff
  2358   fix_least_less       ~> fix_least_below
  2359   dist_less_one        ~> dist_below_one
  2360   less_ONE             ~> below_ONE
  2361   ONE_less_iff         ~> ONE_below_iff
  2362   less_sinlD           ~> below_sinlD
  2363   less_sinrD           ~> below_sinrD
  2364 
  2365 
  2366 *** FOL and ZF ***
  2367 
  2368 * All constant names are now qualified internally and use proper
  2369 identifiers, e.g. "IFOL.eq" instead of "op =".  INCOMPATIBILITY.
  2370 
  2371 
  2372 *** ML ***
  2373 
  2374 * Antiquotation @{assert} inlines a function bool -> unit that raises
  2375 Fail if the argument is false.  Due to inlining the source position of
  2376 failed assertions is included in the error output.
  2377 
  2378 * Discontinued antiquotation @{theory_ref}, which is obsolete since ML
  2379 text is in practice always evaluated with a stable theory checkpoint.
  2380 Minor INCOMPATIBILITY, use (Theory.check_thy @{theory}) instead.
  2381 
  2382 * Antiquotation @{theory A} refers to theory A from the ancestry of
  2383 the current context, not any accidental theory loader state as before.
  2384 Potential INCOMPATIBILITY, subtle change in semantics.
  2385 
  2386 * Syntax.pretty_priority (default 0) configures the required priority
  2387 of pretty-printed output and thus affects insertion of parentheses.
  2388 
  2389 * Syntax.default_root (default "any") configures the inner syntax
  2390 category (nonterminal symbol) for parsing of terms.
  2391 
  2392 * Former exception Library.UnequalLengths now coincides with
  2393 ListPair.UnequalLengths.
  2394 
  2395 * Renamed structure MetaSimplifier to Raw_Simplifier.  Note that the
  2396 main functionality is provided by structure Simplifier.
  2397 
  2398 * Renamed raw "explode" function to "raw_explode" to emphasize its
  2399 meaning.  Note that internally to Isabelle, Symbol.explode is used in
  2400 almost all situations.
  2401 
  2402 * Discontinued obsolete function sys_error and exception SYS_ERROR.
  2403 See implementation manual for further details on exceptions in
  2404 Isabelle/ML.
  2405 
  2406 * Renamed setmp_noncritical to Unsynchronized.setmp to emphasize its
  2407 meaning.
  2408 
  2409 * Renamed structure PureThy to Pure_Thy and moved most of its
  2410 operations to structure Global_Theory, to emphasize that this is
  2411 rarely-used global-only stuff.
  2412 
  2413 * Discontinued Output.debug.  Minor INCOMPATIBILITY, use plain writeln
  2414 instead (or tracing for high-volume output).
  2415 
  2416 * Configuration option show_question_marks only affects regular pretty
  2417 printing of types and terms, not raw Term.string_of_vname.
  2418 
  2419 * ML_Context.thm and ML_Context.thms are no longer pervasive.  Rare
  2420 INCOMPATIBILITY, superseded by static antiquotations @{thm} and
  2421 @{thms} for most purposes.
  2422 
  2423 * ML structure Unsynchronized is never opened, not even in Isar
  2424 interaction mode as before.  Old Unsynchronized.set etc. have been
  2425 discontinued -- use plain := instead.  This should be *rare* anyway,
  2426 since modern tools always work via official context data, notably
  2427 configuration options.
  2428 
  2429 * Parallel and asynchronous execution requires special care concerning
  2430 interrupts.  Structure Exn provides some convenience functions that
  2431 avoid working directly with raw Interrupt.  User code must not absorb
  2432 interrupts -- intermediate handling (for cleanup etc.) needs to be
  2433 followed by re-raising of the original exception.  Another common
  2434 source of mistakes are "handle _" patterns, which make the meaning of
  2435 the program subject to physical effects of the environment.
  2436 
  2437 
  2438 
  2439 New in Isabelle2009-2 (June 2010)
  2440 ---------------------------------
  2441 
  2442 *** General ***
  2443 
  2444 * Authentic syntax for *all* logical entities (type classes, type
  2445 constructors, term constants): provides simple and robust
  2446 correspondence between formal entities and concrete syntax.  Within
  2447 the parse tree / AST representations, "constants" are decorated by
  2448 their category (class, type, const) and spelled out explicitly with
  2449 their full internal name.
  2450 
  2451 Substantial INCOMPATIBILITY concerning low-level syntax declarations
  2452 and translations (translation rules and translation functions in ML).
  2453 Some hints on upgrading:
  2454 
  2455   - Many existing uses of 'syntax' and 'translations' can be replaced
  2456     by more modern 'type_notation', 'notation' and 'abbreviation',
  2457     which are independent of this issue.
  2458 
  2459   - 'translations' require markup within the AST; the term syntax
  2460     provides the following special forms:
  2461 
  2462       CONST c   -- produces syntax version of constant c from context
  2463       XCONST c  -- literally c, checked as constant from context
  2464       c         -- literally c, if declared by 'syntax'
  2465 
  2466     Plain identifiers are treated as AST variables -- occasionally the
  2467     system indicates accidental variables via the error "rhs contains
  2468     extra variables".
  2469 
  2470     Type classes and type constructors are marked according to their
  2471     concrete syntax.  Some old translations rules need to be written
  2472     for the "type" category, using type constructor application
  2473     instead of pseudo-term application of the default category
  2474     "logic".
  2475 
  2476   - 'parse_translation' etc. in ML may use the following
  2477     antiquotations:
  2478 
  2479       @{class_syntax c}   -- type class c within parse tree / AST
  2480       @{term_syntax c}    -- type constructor c within parse tree / AST
  2481       @{const_syntax c}   -- ML version of "CONST c" above
  2482       @{syntax_const c}   -- literally c (checked wrt. 'syntax' declarations)
  2483 
  2484   - Literal types within 'typed_print_translations', i.e. those *not*
  2485     represented as pseudo-terms are represented verbatim.  Use @{class
  2486     c} or @{type_name c} here instead of the above syntax
  2487     antiquotations.
  2488 
  2489 Note that old non-authentic syntax was based on unqualified base
  2490 names, so all of the above "constant" names would coincide.  Recall
  2491 that 'print_syntax' and ML_command "set Syntax.trace_ast" help to
  2492 diagnose syntax problems.
  2493 
  2494 * Type constructors admit general mixfix syntax, not just infix.
  2495 
  2496 * Concrete syntax may be attached to local entities without a proof
  2497 body, too.  This works via regular mixfix annotations for 'fix',
  2498 'def', 'obtain' etc. or via the explicit 'write' command, which is
  2499 similar to the 'notation' command in theory specifications.
  2500 
  2501 * Discontinued unnamed infix syntax (legacy feature for many years) --
  2502 need to specify constant name and syntax separately.  Internal ML
  2503 datatype constructors have been renamed from InfixName to Infix etc.
  2504 Minor INCOMPATIBILITY.
  2505 
  2506 * Schematic theorem statements need to be explicitly markup as such,
  2507 via commands 'schematic_lemma', 'schematic_theorem',
  2508 'schematic_corollary'.  Thus the relevance of the proof is made
  2509 syntactically clear, which impacts performance in a parallel or
  2510 asynchronous interactive environment.  Minor INCOMPATIBILITY.
  2511 
  2512 * Use of cumulative prems via "!" in some proof methods has been
  2513 discontinued (old legacy feature).
  2514 
  2515 * References 'trace_simp' and 'debug_simp' have been replaced by
  2516 configuration options stored in the context. Enabling tracing (the
  2517 case of debugging is similar) in proofs works via
  2518 
  2519   using [[trace_simp = true]]
  2520 
  2521 Tracing is then active for all invocations of the simplifier in
  2522 subsequent goal refinement steps. Tracing may also still be enabled or
  2523 disabled via the ProofGeneral settings menu.
  2524 
  2525 * Separate commands 'hide_class', 'hide_type', 'hide_const',
  2526 'hide_fact' replace the former 'hide' KIND command.  Minor
  2527 INCOMPATIBILITY.
  2528 
  2529 * Improved parallelism of proof term normalization: usedir -p2 -q0 is
  2530 more efficient than combinations with -q1 or -q2.
  2531 
  2532 
  2533 *** Pure ***
  2534 
  2535 * Proofterms record type-class reasoning explicitly, using the
  2536 "unconstrain" operation internally.  This eliminates all sort
  2537 constraints from a theorem and proof, introducing explicit
  2538 OFCLASS-premises.  On the proof term level, this operation is
  2539 automatically applied at theorem boundaries, such that closed proofs
  2540 are always free of sort constraints.  INCOMPATIBILITY for tools that
  2541 inspect proof terms.
  2542 
  2543 * Local theory specifications may depend on extra type variables that
  2544 are not present in the result type -- arguments TYPE('a) :: 'a itself
  2545 are added internally.  For example:
  2546 
  2547   definition unitary :: bool where "unitary = (ALL (x::'a) y. x = y)"
  2548 
  2549 * Predicates of locales introduced by classes carry a mandatory
  2550 "class" prefix.  INCOMPATIBILITY.
  2551 
  2552 * Vacuous class specifications observe default sort.  INCOMPATIBILITY.
  2553 
  2554 * Old 'axclass' command has been discontinued.  INCOMPATIBILITY, use
  2555 'class' instead.
  2556 
  2557 * Command 'code_reflect' allows to incorporate generated ML code into
  2558 runtime environment; replaces immature code_datatype antiquotation.
  2559 INCOMPATIBILITY.
  2560 
  2561 * Code generator: simple concept for abstract datatypes obeying
  2562 invariants.
  2563 
  2564 * Code generator: details of internal data cache have no impact on the
  2565 user space functionality any longer.
  2566 
  2567 * Methods "unfold_locales" and "intro_locales" ignore non-locale
  2568 subgoals.  This is more appropriate for interpretations with 'where'.
  2569 INCOMPATIBILITY.
  2570 
  2571 * Command 'example_proof' opens an empty proof body.  This allows to
  2572 experiment with Isar, without producing any persistent result.
  2573 
  2574 * Commands 'type_notation' and 'no_type_notation' declare type syntax
  2575 within a local theory context, with explicit checking of the
  2576 constructors involved (in contrast to the raw 'syntax' versions).
  2577 
  2578 * Commands 'types' and 'typedecl' now work within a local theory
  2579 context -- without introducing dependencies on parameters or
  2580 assumptions, which is not possible in Isabelle/Pure.
  2581 
  2582 * Command 'defaultsort' has been renamed to 'default_sort', it works
  2583 within a local theory context.  Minor INCOMPATIBILITY.
  2584 
  2585 
  2586 *** HOL ***
  2587 
  2588 * Command 'typedef' now works within a local theory context -- without
  2589 introducing dependencies on parameters or assumptions, which is not
  2590 possible in Isabelle/Pure/HOL.  Note that the logical environment may
  2591 contain multiple interpretations of local typedefs (with different
  2592 non-emptiness proofs), even in a global theory context.
  2593 
  2594 * New package for quotient types.  Commands 'quotient_type' and
  2595 'quotient_definition' may be used for defining types and constants by
  2596 quotient constructions.  An example is the type of integers created by
  2597 quotienting pairs of natural numbers:
  2598 
  2599   fun
  2600     intrel :: "(nat * nat) => (nat * nat) => bool"
  2601   where
  2602     "intrel (x, y) (u, v) = (x + v = u + y)"
  2603 
  2604   quotient_type int = "nat * nat" / intrel
  2605     by (auto simp add: equivp_def expand_fun_eq)
  2606 
  2607   quotient_definition
  2608     "0::int" is "(0::nat, 0::nat)"
  2609 
  2610 The method "lifting" can be used to lift of theorems from the
  2611 underlying "raw" type to the quotient type.  The example
  2612 src/HOL/Quotient_Examples/FSet.thy includes such a quotient
  2613 construction and provides a reasoning infrastructure for finite sets.
  2614 
  2615 * Renamed Library/Quotient.thy to Library/Quotient_Type.thy to avoid
  2616 clash with new theory Quotient in Main HOL.
  2617 
  2618 * Moved the SMT binding into the main HOL session, eliminating
  2619 separate HOL-SMT session.
  2620 
  2621 * List membership infix mem operation is only an input abbreviation.
  2622 INCOMPATIBILITY.
  2623 
  2624 * Theory Library/Word.thy has been removed.  Use library Word/Word.thy
  2625 for future developements; former Library/Word.thy is still present in
  2626 the AFP entry RSAPPS.
  2627 
  2628 * Theorem Int.int_induct renamed to Int.int_of_nat_induct and is no
  2629 longer shadowed.  INCOMPATIBILITY.
  2630 
  2631 * Dropped theorem duplicate comp_arith; use semiring_norm instead.
  2632 INCOMPATIBILITY.
  2633 
  2634 * Dropped theorem RealPow.real_sq_order; use power2_le_imp_le instead.
  2635 INCOMPATIBILITY.
  2636 
  2637 * Dropped normalizing_semiring etc; use the facts in semiring classes
  2638 instead.  INCOMPATIBILITY.
  2639 
  2640 * Dropped several real-specific versions of lemmas about floor and
  2641 ceiling; use the generic lemmas from theory "Archimedean_Field"
  2642 instead.  INCOMPATIBILITY.
  2643 
  2644   floor_number_of_eq         ~> floor_number_of
  2645   le_floor_eq_number_of      ~> number_of_le_floor
  2646   le_floor_eq_zero           ~> zero_le_floor
  2647   le_floor_eq_one            ~> one_le_floor
  2648   floor_less_eq_number_of    ~> floor_less_number_of
  2649   floor_less_eq_zero         ~> floor_less_zero
  2650   floor_less_eq_one          ~> floor_less_one
  2651   less_floor_eq_number_of    ~> number_of_less_floor
  2652   less_floor_eq_zero         ~> zero_less_floor
  2653   less_floor_eq_one          ~> one_less_floor
  2654   floor_le_eq_number_of      ~> floor_le_number_of
  2655   floor_le_eq_zero           ~> floor_le_zero
  2656   floor_le_eq_one            ~> floor_le_one
  2657   floor_subtract_number_of   ~> floor_diff_number_of
  2658   floor_subtract_one         ~> floor_diff_one
  2659   ceiling_number_of_eq       ~> ceiling_number_of
  2660   ceiling_le_eq_number_of    ~> ceiling_le_number_of
  2661   ceiling_le_zero_eq         ~> ceiling_le_zero
  2662   ceiling_le_eq_one          ~> ceiling_le_one
  2663   less_ceiling_eq_number_of  ~> number_of_less_ceiling
  2664   less_ceiling_eq_zero       ~> zero_less_ceiling
  2665   less_ceiling_eq_one        ~> one_less_ceiling
  2666   ceiling_less_eq_number_of  ~> ceiling_less_number_of
  2667   ceiling_less_eq_zero       ~> ceiling_less_zero
  2668   ceiling_less_eq_one        ~> ceiling_less_one
  2669   le_ceiling_eq_number_of    ~> number_of_le_ceiling
  2670   le_ceiling_eq_zero         ~> zero_le_ceiling
  2671   le_ceiling_eq_one          ~> one_le_ceiling
  2672   ceiling_subtract_number_of ~> ceiling_diff_number_of
  2673   ceiling_subtract_one       ~> ceiling_diff_one
  2674 
  2675 * Theory "Finite_Set": various folding_XXX locales facilitate the
  2676 application of the various fold combinators on finite sets.
  2677 
  2678 * Library theory "RBT" renamed to "RBT_Impl"; new library theory "RBT"
  2679 provides abstract red-black tree type which is backed by "RBT_Impl" as
  2680 implementation.  INCOMPATIBILTY.
  2681 
  2682 * Theory Library/Coinductive_List has been removed -- superseded by
  2683 AFP/thys/Coinductive.
  2684 
  2685 * Theory PReal, including the type "preal" and related operations, has
  2686 been removed.  INCOMPATIBILITY.
  2687 
  2688 * Real: new development using Cauchy Sequences.
  2689 
  2690 * Split off theory "Big_Operators" containing setsum, setprod,
  2691 Inf_fin, Sup_fin, Min, Max from theory Finite_Set.  INCOMPATIBILITY.
  2692 
  2693 * Theory "Rational" renamed to "Rat", for consistency with "Nat",
  2694 "Int" etc.  INCOMPATIBILITY.
  2695 
  2696 * Constant Rat.normalize needs to be qualified.  INCOMPATIBILITY.
  2697 
  2698 * New set of rules "ac_simps" provides combined assoc / commute
  2699 rewrites for all interpretations of the appropriate generic locales.
  2700 
  2701 * Renamed theory "OrderedGroup" to "Groups" and split theory
  2702 "Ring_and_Field" into theories "Rings" and "Fields"; for more
  2703 appropriate and more consistent names suitable for name prefixes
  2704 within the HOL theories.  INCOMPATIBILITY.
  2705 
  2706 * Some generic constants have been put to appropriate theories:
  2707   - less_eq, less: Orderings
  2708   - zero, one, plus, minus, uminus, times, abs, sgn: Groups
  2709   - inverse, divide: Rings
  2710 INCOMPATIBILITY.
  2711 
  2712 * More consistent naming of type classes involving orderings (and
  2713 lattices):
  2714 
  2715     lower_semilattice                   ~> semilattice_inf
  2716     upper_semilattice                   ~> semilattice_sup
  2717 
  2718     dense_linear_order                  ~> dense_linorder
  2719 
  2720     pordered_ab_group_add               ~> ordered_ab_group_add
  2721     pordered_ab_group_add_abs           ~> ordered_ab_group_add_abs
  2722     pordered_ab_semigroup_add           ~> ordered_ab_semigroup_add
  2723     pordered_ab_semigroup_add_imp_le    ~> ordered_ab_semigroup_add_imp_le
  2724     pordered_cancel_ab_semigroup_add    ~> ordered_cancel_ab_semigroup_add
  2725     pordered_cancel_comm_semiring       ~> ordered_cancel_comm_semiring
  2726     pordered_cancel_semiring            ~> ordered_cancel_semiring
  2727     pordered_comm_monoid_add            ~> ordered_comm_monoid_add
  2728     pordered_comm_ring                  ~> ordered_comm_ring
  2729     pordered_comm_semiring              ~> ordered_comm_semiring
  2730     pordered_ring                       ~> ordered_ring
  2731     pordered_ring_abs                   ~> ordered_ring_abs
  2732     pordered_semiring                   ~> ordered_semiring
  2733 
  2734     ordered_ab_group_add                ~> linordered_ab_group_add
  2735     ordered_ab_semigroup_add            ~> linordered_ab_semigroup_add
  2736     ordered_cancel_ab_semigroup_add     ~> linordered_cancel_ab_semigroup_add
  2737     ordered_comm_semiring_strict        ~> linordered_comm_semiring_strict
  2738     ordered_field                       ~> linordered_field
  2739     ordered_field_no_lb                 ~> linordered_field_no_lb
  2740     ordered_field_no_ub                 ~> linordered_field_no_ub
  2741     ordered_field_dense_linear_order    ~> dense_linordered_field
  2742     ordered_idom                        ~> linordered_idom
  2743     ordered_ring                        ~> linordered_ring
  2744     ordered_ring_le_cancel_factor       ~> linordered_ring_le_cancel_factor
  2745     ordered_ring_less_cancel_factor     ~> linordered_ring_less_cancel_factor
  2746     ordered_ring_strict                 ~> linordered_ring_strict
  2747     ordered_semidom                     ~> linordered_semidom
  2748     ordered_semiring                    ~> linordered_semiring
  2749     ordered_semiring_1                  ~> linordered_semiring_1
  2750     ordered_semiring_1_strict           ~> linordered_semiring_1_strict
  2751     ordered_semiring_strict             ~> linordered_semiring_strict
  2752 
  2753   The following slightly odd type classes have been moved to a
  2754   separate theory Library/Lattice_Algebras:
  2755 
  2756     lordered_ab_group_add               ~> lattice_ab_group_add
  2757     lordered_ab_group_add_abs           ~> lattice_ab_group_add_abs
  2758     lordered_ab_group_add_meet          ~> semilattice_inf_ab_group_add
  2759     lordered_ab_group_add_join          ~> semilattice_sup_ab_group_add
  2760     lordered_ring                       ~> lattice_ring
  2761 
  2762 INCOMPATIBILITY.
  2763 
  2764 * Refined field classes:
  2765   - classes division_ring_inverse_zero, field_inverse_zero,
  2766     linordered_field_inverse_zero include rule inverse 0 = 0 --
  2767     subsumes former division_by_zero class;
  2768   - numerous lemmas have been ported from field to division_ring.
  2769 INCOMPATIBILITY.
  2770 
  2771 * Refined algebra theorem collections:
  2772   - dropped theorem group group_simps, use algebra_simps instead;
  2773   - dropped theorem group ring_simps, use field_simps instead;
  2774   - proper theorem collection field_simps subsumes former theorem
  2775     groups field_eq_simps and field_simps;
  2776   - dropped lemma eq_minus_self_iff which is a duplicate for
  2777     equal_neg_zero.
  2778 INCOMPATIBILITY.
  2779 
  2780 * Theory Finite_Set and List: some lemmas have been generalized from
  2781 sets to lattices:
  2782 
  2783   fun_left_comm_idem_inter      ~> fun_left_comm_idem_inf
  2784   fun_left_comm_idem_union      ~> fun_left_comm_idem_sup
  2785   inter_Inter_fold_inter        ~> inf_Inf_fold_inf
  2786   union_Union_fold_union        ~> sup_Sup_fold_sup
  2787   Inter_fold_inter              ~> Inf_fold_inf
  2788   Union_fold_union              ~> Sup_fold_sup
  2789   inter_INTER_fold_inter        ~> inf_INFI_fold_inf
  2790   union_UNION_fold_union        ~> sup_SUPR_fold_sup
  2791   INTER_fold_inter              ~> INFI_fold_inf
  2792   UNION_fold_union              ~> SUPR_fold_sup
  2793 
  2794 * Theory "Complete_Lattice": lemmas top_def and bot_def have been
  2795 replaced by the more convenient lemmas Inf_empty and Sup_empty.
  2796 Dropped lemmas Inf_insert_simp and Sup_insert_simp, which are subsumed
  2797 by Inf_insert and Sup_insert.  Lemmas Inf_UNIV and Sup_UNIV replace
  2798 former Inf_Univ and Sup_Univ.  Lemmas inf_top_right and sup_bot_right
  2799 subsume inf_top and sup_bot respectively.  INCOMPATIBILITY.
  2800 
  2801 * Reorganized theory Multiset: swapped notation of pointwise and
  2802 multiset order:
  2803 
  2804   - pointwise ordering is instance of class order with standard syntax
  2805     <= and <;
  2806   - multiset ordering has syntax <=# and <#; partial order properties
  2807     are provided by means of interpretation with prefix
  2808     multiset_order;
  2809   - less duplication, less historical organization of sections,
  2810     conversion from associations lists to multisets, rudimentary code
  2811     generation;
  2812   - use insert_DiffM2 [symmetric] instead of elem_imp_eq_diff_union,
  2813     if needed.
  2814 
  2815 Renamed:
  2816 
  2817   multiset_eq_conv_count_eq  ~>  multiset_ext_iff
  2818   multi_count_ext  ~>  multiset_ext
  2819   diff_union_inverse2  ~>  diff_union_cancelR
  2820 
  2821 INCOMPATIBILITY.
  2822 
  2823 * Theory Permutation: replaced local "remove" by List.remove1.
  2824 
  2825 * Code generation: ML and OCaml code is decorated with signatures.
  2826 
  2827 * Theory List: added transpose.
  2828 
  2829 * Library/Nat_Bijection.thy is a collection of bijective functions
  2830 between nat and other types, which supersedes the older libraries
  2831 Library/Nat_Int_Bij.thy and HOLCF/NatIso.thy.  INCOMPATIBILITY.
  2832 
  2833   Constants:
  2834   Nat_Int_Bij.nat2_to_nat         ~> prod_encode
  2835   Nat_Int_Bij.nat_to_nat2         ~> prod_decode
  2836   Nat_Int_Bij.int_to_nat_bij      ~> int_encode
  2837   Nat_Int_Bij.nat_to_int_bij      ~> int_decode
  2838   Countable.pair_encode           ~> prod_encode
  2839   NatIso.prod2nat                 ~> prod_encode
  2840   NatIso.nat2prod                 ~> prod_decode
  2841   NatIso.sum2nat                  ~> sum_encode
  2842   NatIso.nat2sum                  ~> sum_decode
  2843   NatIso.list2nat                 ~> list_encode
  2844   NatIso.nat2list                 ~> list_decode
  2845   NatIso.set2nat                  ~> set_encode
  2846   NatIso.nat2set                  ~> set_decode
  2847 
  2848   Lemmas:
  2849   Nat_Int_Bij.bij_nat_to_int_bij  ~> bij_int_decode
  2850   Nat_Int_Bij.nat2_to_nat_inj     ~> inj_prod_encode
  2851   Nat_Int_Bij.nat2_to_nat_surj    ~> surj_prod_encode
  2852   Nat_Int_Bij.nat_to_nat2_inj     ~> inj_prod_decode
  2853   Nat_Int_Bij.nat_to_nat2_surj    ~> surj_prod_decode
  2854   Nat_Int_Bij.i2n_n2i_id          ~> int_encode_inverse
  2855   Nat_Int_Bij.n2i_i2n_id          ~> int_decode_inverse
  2856   Nat_Int_Bij.surj_nat_to_int_bij ~> surj_int_encode
  2857   Nat_Int_Bij.surj_int_to_nat_bij ~> surj_int_decode
  2858   Nat_Int_Bij.inj_nat_to_int_bij  ~> inj_int_encode
  2859   Nat_Int_Bij.inj_int_to_nat_bij  ~> inj_int_decode
  2860   Nat_Int_Bij.bij_nat_to_int_bij  ~> bij_int_encode
  2861   Nat_Int_Bij.bij_int_to_nat_bij  ~> bij_int_decode
  2862 
  2863 * Sledgehammer:
  2864   - Renamed ATP commands:
  2865     atp_info     ~> sledgehammer running_atps
  2866     atp_kill     ~> sledgehammer kill_atps
  2867     atp_messages ~> sledgehammer messages
  2868     atp_minimize ~> sledgehammer minimize
  2869     print_atps   ~> sledgehammer available_atps
  2870     INCOMPATIBILITY.
  2871   - Added user's manual ("isabelle doc sledgehammer").
  2872   - Added option syntax and "sledgehammer_params" to customize
  2873     Sledgehammer's behavior.  See the manual for details.
  2874   - Modified the Isar proof reconstruction code so that it produces
  2875     direct proofs rather than proofs by contradiction.  (This feature
  2876     is still experimental.)
  2877   - Made Isar proof reconstruction work for SPASS, remote ATPs, and in
  2878     full-typed mode.
  2879   - Added support for TPTP syntax for SPASS via the "spass_tptp" ATP.
  2880 
  2881 * Nitpick:
  2882   - Added and implemented "binary_ints" and "bits" options.
  2883   - Added "std" option and implemented support for nonstandard models.
  2884   - Added and implemented "finitize" option to improve the precision
  2885     of infinite datatypes based on a monotonicity analysis.
  2886   - Added support for quotient types.
  2887   - Added support for "specification" and "ax_specification"
  2888     constructs.
  2889   - Added support for local definitions (for "function" and
  2890     "termination" proofs).
  2891   - Added support for term postprocessors.
  2892   - Optimized "Multiset.multiset" and "FinFun.finfun".
  2893   - Improved efficiency of "destroy_constrs" optimization.
  2894   - Fixed soundness bugs related to "destroy_constrs" optimization and
  2895     record getters.
  2896   - Fixed soundness bug related to higher-order constructors.
  2897   - Fixed soundness bug when "full_descrs" is enabled.
  2898   - Improved precision of set constructs.
  2899   - Added "atoms" option.
  2900   - Added cache to speed up repeated Kodkod invocations on the same
  2901     problems.
  2902   - Renamed "MiniSatJNI", "zChaffJNI", "BerkMinAlloy", and
  2903     "SAT4JLight" to "MiniSat_JNI", "zChaff_JNI", "BerkMin_Alloy", and
  2904     "SAT4J_Light".  INCOMPATIBILITY.
  2905   - Removed "skolemize", "uncurry", "sym_break", "flatten_prop",
  2906     "sharing_depth", and "show_skolems" options.  INCOMPATIBILITY.
  2907   - Removed "nitpick_intro" attribute.  INCOMPATIBILITY.
  2908 
  2909 * Method "induct" now takes instantiations of the form t, where t is not
  2910   a variable, as a shorthand for "x == t", where x is a fresh variable.
  2911   If this is not intended, t has to be enclosed in parentheses.
  2912   By default, the equalities generated by definitional instantiations
  2913   are pre-simplified, which may cause parameters of inductive cases
  2914   to disappear, or may even delete some of the inductive cases.
  2915   Use "induct (no_simp)" instead of "induct" to restore the old
  2916   behaviour. The (no_simp) option is also understood by the "cases"
  2917   and "nominal_induct" methods, which now perform pre-simplification, too.
  2918   INCOMPATIBILITY.
  2919 
  2920 
  2921 *** HOLCF ***
  2922 
  2923 * Variable names in lemmas generated by the domain package have
  2924 changed; the naming scheme is now consistent with the HOL datatype
  2925 package.  Some proof scripts may be affected, INCOMPATIBILITY.
  2926 
  2927 * The domain package no longer defines the function "foo_copy" for
  2928 recursive domain "foo".  The reach lemma is now stated directly in
  2929 terms of "foo_take".  Lemmas and proofs that mention "foo_copy" must
  2930 be reformulated in terms of "foo_take", INCOMPATIBILITY.
  2931 
  2932 * Most definedness lemmas generated by the domain package (previously
  2933 of the form "x ~= UU ==> foo$x ~= UU") now have an if-and-only-if form
  2934 like "foo$x = UU <-> x = UU", which works better as a simp rule.
  2935 Proofs that used definedness lemmas as intro rules may break,
  2936 potential INCOMPATIBILITY.
  2937 
  2938 * Induction and casedist rules generated by the domain package now
  2939 declare proper case_names (one called "bottom", and one named for each
  2940 constructor).  INCOMPATIBILITY.
  2941 
  2942 * For mutually-recursive domains, separate "reach" and "take_lemma"
  2943 rules are generated for each domain, INCOMPATIBILITY.
  2944 
  2945   foo_bar.reach       ~> foo.reach  bar.reach
  2946   foo_bar.take_lemmas ~> foo.take_lemma  bar.take_lemma
  2947 
  2948 * Some lemmas generated by the domain package have been renamed for
  2949 consistency with the datatype package, INCOMPATIBILITY.
  2950 
  2951   foo.ind        ~> foo.induct
  2952   foo.finite_ind ~> foo.finite_induct
  2953   foo.coind      ~> foo.coinduct
  2954   foo.casedist   ~> foo.exhaust
  2955   foo.exhaust    ~> foo.nchotomy
  2956 
  2957 * For consistency with other definition packages, the fixrec package
  2958 now generates qualified theorem names, INCOMPATIBILITY.
  2959 
  2960   foo_simps  ~> foo.simps
  2961   foo_unfold ~> foo.unfold
  2962   foo_induct ~> foo.induct
  2963 
  2964 * The "fixrec_simp" attribute has been removed.  The "fixrec_simp"
  2965 method and internal fixrec proofs now use the default simpset instead.
  2966 INCOMPATIBILITY.
  2967 
  2968 * The "contlub" predicate has been removed.  Proof scripts should use
  2969 lemma contI2 in place of monocontlub2cont, INCOMPATIBILITY.
  2970 
  2971 * The "admw" predicate has been removed, INCOMPATIBILITY.
  2972 
  2973 * The constants cpair, cfst, and csnd have been removed in favor of
  2974 Pair, fst, and snd from Isabelle/HOL, INCOMPATIBILITY.
  2975 
  2976 
  2977 *** ML ***
  2978 
  2979 * Antiquotations for basic formal entities:
  2980 
  2981     @{class NAME}         -- type class
  2982     @{class_syntax NAME}  -- syntax representation of the above
  2983 
  2984     @{type_name NAME}     -- logical type
  2985     @{type_abbrev NAME}   -- type abbreviation
  2986     @{nonterminal NAME}   -- type of concrete syntactic category
  2987     @{type_syntax NAME}   -- syntax representation of any of the above
  2988 
  2989     @{const_name NAME}    -- logical constant (INCOMPATIBILITY)
  2990     @{const_abbrev NAME}  -- abbreviated constant
  2991     @{const_syntax NAME}  -- syntax representation of any of the above
  2992 
  2993 * Antiquotation @{syntax_const NAME} ensures that NAME refers to a raw
  2994 syntax constant (cf. 'syntax' command).
  2995 
  2996 * Antiquotation @{make_string} inlines a function to print arbitrary
  2997 values similar to the ML toplevel.  The result is compiler dependent
  2998 and may fall back on "?" in certain situations.
  2999 
  3000 * Diagnostic commands 'ML_val' and 'ML_command' may refer to
  3001 antiquotations @{Isar.state} and @{Isar.goal}.  This replaces impure
  3002 Isar.state() and Isar.goal(), which belong to the old TTY loop and do
  3003 not work with the asynchronous Isar document model.
  3004 
  3005 * Configuration options now admit dynamic default values, depending on
  3006 the context or even global references.
  3007 
  3008 * SHA1.digest digests strings according to SHA-1 (see RFC 3174).  It
  3009 uses an efficient external library if available (for Poly/ML).
  3010 
  3011 * Renamed some important ML structures, while keeping the old names
  3012 for some time as aliases within the structure Legacy:
  3013 
  3014   OuterKeyword  ~>  Keyword
  3015   OuterLex      ~>  Token
  3016   OuterParse    ~>  Parse
  3017   OuterSyntax   ~>  Outer_Syntax
  3018   PrintMode     ~>  Print_Mode
  3019   SpecParse     ~>  Parse_Spec
  3020   ThyInfo       ~>  Thy_Info
  3021   ThyLoad       ~>  Thy_Load
  3022   ThyOutput     ~>  Thy_Output
  3023   TypeInfer     ~>  Type_Infer
  3024 
  3025 Note that "open Legacy" simplifies porting of sources, but forgetting
  3026 to remove it again will complicate porting again in the future.
  3027 
  3028 * Most operations that refer to a global context are named
  3029 accordingly, e.g. Simplifier.global_context or
  3030 ProofContext.init_global.  There are some situations where a global
  3031 context actually works, but under normal circumstances one needs to
  3032 pass the proper local context through the code!
  3033 
  3034 * Discontinued old TheoryDataFun with its copy/init operation -- data
  3035 needs to be pure.  Functor Theory_Data_PP retains the traditional
  3036 Pretty.pp argument to merge, which is absent in the standard
  3037 Theory_Data version.
  3038 
  3039 * Sorts.certify_sort and derived "cert" operations for types and terms
  3040 no longer minimize sorts.  Thus certification at the boundary of the
  3041 inference kernel becomes invariant under addition of class relations,
  3042 which is an important monotonicity principle.  Sorts are now minimized
  3043 in the syntax layer only, at the boundary between the end-user and the
  3044 system.  Subtle INCOMPATIBILITY, may have to use Sign.minimize_sort
  3045 explicitly in rare situations.
  3046 
  3047 * Renamed old-style Drule.standard to Drule.export_without_context, to
  3048 emphasize that this is in no way a standard operation.
  3049 INCOMPATIBILITY.
  3050 
  3051 * Subgoal.FOCUS (and variants): resulting goal state is normalized as
  3052 usual for resolution.  Rare INCOMPATIBILITY.
  3053 
  3054 * Renamed varify/unvarify operations to varify_global/unvarify_global
  3055 to emphasize that these only work in a global situation (which is
  3056 quite rare).
  3057 
  3058 * Curried take and drop in library.ML; negative length is interpreted
  3059 as infinity (as in chop).  Subtle INCOMPATIBILITY.
  3060 
  3061 * Proof terms: type substitutions on proof constants now use canonical
  3062 order of type variables.  INCOMPATIBILITY for tools working with proof
  3063 terms.
  3064 
  3065 * Raw axioms/defs may no longer carry sort constraints, and raw defs
  3066 may no longer carry premises.  User-level specifications are
  3067 transformed accordingly by Thm.add_axiom/add_def.
  3068 
  3069 
  3070 *** System ***
  3071 
  3072 * Discontinued special HOL_USEDIR_OPTIONS for the main HOL image;
  3073 ISABELLE_USEDIR_OPTIONS applies uniformly to all sessions.  Note that
  3074 proof terms are enabled unconditionally in the new HOL-Proofs image.
  3075 
  3076 * Discontinued old ISABELLE and ISATOOL environment settings (legacy
  3077 feature since Isabelle2009).  Use ISABELLE_PROCESS and ISABELLE_TOOL,
  3078 respectively.
  3079 
  3080 * Old lib/scripts/polyml-platform is superseded by the
  3081 ISABELLE_PLATFORM setting variable, which defaults to the 32 bit
  3082 variant, even on a 64 bit machine.  The following example setting
  3083 prefers 64 bit if available:
  3084 
  3085   ML_PLATFORM="${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM}"
  3086 
  3087 * The preliminary Isabelle/jEdit application demonstrates the emerging
  3088 Isabelle/Scala layer for advanced prover interaction and integration.
  3089 See src/Tools/jEdit or "isabelle jedit" provided by the properly built
  3090 component.
  3091 
  3092 * "IsabelleText" is a Unicode font derived from Bitstream Vera Mono
  3093 and Bluesky TeX fonts.  It provides the usual Isabelle symbols,
  3094 similar to the default assignment of the document preparation system
  3095 (cf. isabellesym.sty).  The Isabelle/Scala class Isabelle_System
  3096 provides some operations for direct access to the font without asking
  3097 the user for manual installation.
  3098 
  3099 
  3100 
  3101 New in Isabelle2009-1 (December 2009)
  3102 -------------------------------------
  3103 
  3104 *** General ***
  3105 
  3106 * Discontinued old form of "escaped symbols" such as \\<forall>.  Only
  3107 one backslash should be used, even in ML sources.
  3108 
  3109 
  3110 *** Pure ***
  3111 
  3112 * Locale interpretation propagates mixins along the locale hierarchy.
  3113 The currently only available mixins are the equations used to map
  3114 local definitions to terms of the target domain of an interpretation.
  3115 
  3116 * Reactivated diagnostic command 'print_interps'.  Use "print_interps
  3117 loc" to print all interpretations of locale "loc" in the theory.
  3118 Interpretations in proofs are not shown.
  3119 
  3120 * Thoroughly revised locales tutorial.  New section on conditional
  3121 interpretation.
  3122 
  3123 * On instantiation of classes, remaining undefined class parameters
  3124 are formally declared.  INCOMPATIBILITY.
  3125 
  3126 
  3127 *** Document preparation ***
  3128 
  3129 * New generalized style concept for printing terms: @{foo (style) ...}
  3130 instead of @{foo_style style ...}  (old form is still retained for
  3131 backward compatibility).  Styles can be also applied for
  3132 antiquotations prop, term_type and typeof.
  3133 
  3134 
  3135 *** HOL ***
  3136 
  3137 * New proof method "smt" for a combination of first-order logic with
  3138 equality, linear and nonlinear (natural/integer/real) arithmetic, and
  3139 fixed-size bitvectors; there is also basic support for higher-order
  3140 features (esp. lambda abstractions).  It is an incomplete decision
  3141 procedure based on external SMT solvers using the oracle mechanism;
  3142 for the SMT solver Z3, this method is proof-producing.  Certificates
  3143 are provided to avoid calling the external solvers solely for
  3144 re-checking proofs.  Due to a remote SMT service there is no need for
  3145 installing SMT solvers locally.  See src/HOL/SMT.
  3146 
  3147 * New commands to load and prove verification conditions generated by
  3148 the Boogie program verifier or derived systems (e.g. the Verifying C
  3149 Compiler (VCC) or Spec#).  See src/HOL/Boogie.
  3150 
  3151 * New counterexample generator tool 'nitpick' based on the Kodkod
  3152 relational model finder.  See src/HOL/Tools/Nitpick and
  3153 src/HOL/Nitpick_Examples.
  3154 
  3155 * New commands 'code_pred' and 'values' to invoke the predicate
  3156 compiler and to enumerate values of inductive predicates.
  3157 
  3158 * A tabled implementation of the reflexive transitive closure.
  3159 
  3160 * New implementation of quickcheck uses generic code generator;
  3161 default generators are provided for all suitable HOL types, records
  3162 and datatypes.  Old quickcheck can be re-activated importing theory
  3163 Library/SML_Quickcheck.
  3164 
  3165 * New testing tool Mirabelle for automated proof tools.  Applies
  3166 several tools and tactics like sledgehammer, metis, or quickcheck, to
  3167 every proof step in a theory.  To be used in batch mode via the
  3168 "mirabelle" utility.
  3169 
  3170 * New proof method "sos" (sum of squares) for nonlinear real
  3171 arithmetic (originally due to John Harison). It requires theory
  3172 Library/Sum_Of_Squares.  It is not a complete decision procedure but
  3173 works well in practice on quantifier-free real arithmetic with +, -,
  3174 *, ^, =, <= and <, i.e. boolean combinations of equalities and
  3175 inequalities between polynomials.  It makes use of external
  3176 semidefinite programming solvers.  Method "sos" generates a
  3177 certificate that can be pasted into the proof thus avoiding the need
  3178 to call an external tool every time the proof is checked.  See
  3179 src/HOL/Library/Sum_Of_Squares.
  3180 
  3181 * New method "linarith" invokes existing linear arithmetic decision
  3182 procedure only.
  3183 
  3184 * New command 'atp_minimal' reduces result produced by Sledgehammer.
  3185 
  3186 * New Sledgehammer option "Full Types" in Proof General settings menu.
  3187 Causes full type information to be output to the ATPs.  This slows
  3188 ATPs down considerably but eliminates a source of unsound "proofs"
  3189 that fail later.
  3190 
  3191 * New method "metisFT": A version of metis that uses full type
  3192 information in order to avoid failures of proof reconstruction.
  3193 
  3194 * New evaluator "approximate" approximates an real valued term using
  3195 the same method as the approximation method.
  3196 
  3197 * Method "approximate" now supports arithmetic expressions as
  3198 boundaries of intervals and implements interval splitting and Taylor
  3199 series expansion.
  3200 
  3201 * ML antiquotation @{code_datatype} inserts definition of a datatype
  3202 generated by the code generator; e.g. see src/HOL/Predicate.thy.
  3203 
  3204 * New theory SupInf of the supremum and infimum operators for sets of
  3205 reals.
  3206 
  3207 * New theory Probability, which contains a development of measure
  3208 theory, eventually leading to Lebesgue integration and probability.
  3209 
  3210 * Extended Multivariate Analysis to include derivation and Brouwer's
  3211 fixpoint theorem.
  3212 
  3213 * Reorganization of number theory, INCOMPATIBILITY:
  3214   - new number theory development for nat and int, in theories Divides
  3215     and GCD as well as in new session Number_Theory
  3216   - some constants and facts now suffixed with _nat and _int
  3217     accordingly
  3218   - former session NumberTheory now named Old_Number_Theory, including
  3219     theories Legacy_GCD and Primes (prefer Number_Theory if possible)
  3220   - moved theory Pocklington from src/HOL/Library to
  3221     src/HOL/Old_Number_Theory
  3222 
  3223 * Theory GCD includes functions Gcd/GCD and Lcm/LCM for the gcd and
  3224 lcm of finite and infinite sets. It is shown that they form a complete
  3225 lattice.
  3226 
  3227 * Class semiring_div requires superclass no_zero_divisors and proof of
  3228 div_mult_mult1; theorems div_mult_mult1, div_mult_mult2,
  3229 div_mult_mult1_if, div_mult_mult1 and div_mult_mult2 have been
  3230 generalized to class semiring_div, subsuming former theorems
  3231 zdiv_zmult_zmult1, zdiv_zmult_zmult1_if, zdiv_zmult_zmult1 and
  3232 zdiv_zmult_zmult2.  div_mult_mult1 is now [simp] by default.
  3233 INCOMPATIBILITY.
  3234 
  3235 * Refinements to lattice classes and sets:
  3236   - less default intro/elim rules in locale variant, more default
  3237     intro/elim rules in class variant: more uniformity
  3238   - lemma ge_sup_conv renamed to le_sup_iff, in accordance with
  3239     le_inf_iff
  3240   - dropped lemma alias inf_ACI for inf_aci (same for sup_ACI and
  3241     sup_aci)
  3242   - renamed ACI to inf_sup_aci
  3243   - new class "boolean_algebra"
  3244   - class "complete_lattice" moved to separate theory
  3245     "Complete_Lattice"; corresponding constants (and abbreviations)
  3246     renamed and with authentic syntax:
  3247     Set.Inf ~>    Complete_Lattice.Inf
  3248     Set.Sup ~>    Complete_Lattice.Sup
  3249     Set.INFI ~>   Complete_Lattice.INFI
  3250     Set.SUPR ~>   Complete_Lattice.SUPR
  3251     Set.Inter ~>  Complete_Lattice.Inter
  3252     Set.Union ~>  Complete_Lattice.Union
  3253     Set.INTER ~>  Complete_Lattice.INTER
  3254     Set.UNION ~>  Complete_Lattice.UNION
  3255   - authentic syntax for
  3256     Set.Pow
  3257     Set.image
  3258   - mere abbreviations:
  3259     Set.empty               (for bot)
  3260     Set.UNIV                (for top)
  3261     Set.inter               (for inf, formerly Set.Int)
  3262     Set.union               (for sup, formerly Set.Un)
  3263     Complete_Lattice.Inter  (for Inf)
  3264     Complete_Lattice.Union  (for Sup)
  3265     Complete_Lattice.INTER  (for INFI)
  3266     Complete_Lattice.UNION  (for SUPR)
  3267   - object-logic definitions as far as appropriate
  3268 
  3269 INCOMPATIBILITY.  Care is required when theorems Int_subset_iff or
  3270 Un_subset_iff are explicitly deleted as default simp rules; then also
  3271 their lattice counterparts le_inf_iff and le_sup_iff have to be
  3272 deleted to achieve the desired effect.
  3273 
  3274 * Rules inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp
  3275 rules by default any longer; the same applies to min_max.inf_absorb1
  3276 etc.  INCOMPATIBILITY.
  3277 
  3278 * Rules sup_Int_eq and sup_Un_eq are no longer declared as
  3279 pred_set_conv by default.  INCOMPATIBILITY.
  3280 
  3281 * Power operations on relations and functions are now one dedicated
  3282 constant "compow" with infix syntax "^^".  Power operation on
  3283 multiplicative monoids retains syntax "^" and is now defined generic
  3284 in class power.  INCOMPATIBILITY.
  3285 
  3286 * Relation composition "R O S" now has a more standard argument order:
  3287 "R O S = {(x, z). EX y. (x, y) : R & (y, z) : S}".  INCOMPATIBILITY,
  3288 rewrite propositions with "S O R" --> "R O S". Proofs may occasionally
  3289 break, since the O_assoc rule was not rewritten like this.  Fix using
  3290 O_assoc[symmetric].  The same applies to the curried version "R OO S".
  3291 
  3292 * Function "Inv" is renamed to "inv_into" and function "inv" is now an
  3293 abbreviation for "inv_into UNIV".  Lemmas are renamed accordingly.
  3294 INCOMPATIBILITY.
  3295 
  3296 * Most rules produced by inductive and datatype package have mandatory
  3297 prefixes.  INCOMPATIBILITY.
  3298 
  3299 * Changed "DERIV_intros" to a dynamic fact, which can be augmented by
  3300 the attribute of the same name.  Each of the theorems in the list
  3301 DERIV_intros assumes composition with an additional function and
  3302 matches a variable to the derivative, which has to be solved by the
  3303 Simplifier.  Hence (auto intro!: DERIV_intros) computes the derivative
  3304 of most elementary terms.  Former Maclauren.DERIV_tac and
  3305 Maclauren.deriv_tac should be replaced by (auto intro!: DERIV_intros).
  3306 INCOMPATIBILITY.
  3307 
  3308 * Code generator attributes follow the usual underscore convention:
  3309     code_unfold     replaces    code unfold
  3310     code_post       replaces    code post
  3311     etc.
  3312   INCOMPATIBILITY.
  3313 
  3314 * Renamed methods:
  3315     sizechange -> size_change
  3316     induct_scheme -> induction_schema
  3317   INCOMPATIBILITY.
  3318 
  3319 * Discontinued abbreviation "arbitrary" of constant "undefined".
  3320 INCOMPATIBILITY, use "undefined" directly.
  3321 
  3322 * Renamed theorems:
  3323     Suc_eq_add_numeral_1 -> Suc_eq_plus1
  3324     Suc_eq_add_numeral_1_left -> Suc_eq_plus1_left
  3325     Suc_plus1 -> Suc_eq_plus1
  3326     *anti_sym -> *antisym*
  3327     vector_less_eq_def -> vector_le_def
  3328   INCOMPATIBILITY.
  3329 
  3330 * Added theorem List.map_map as [simp].  Removed List.map_compose.
  3331 INCOMPATIBILITY.
  3332 
  3333 * Removed predicate "M hassize n" (<--> card M = n & finite M).
  3334 INCOMPATIBILITY.
  3335 
  3336 
  3337 *** HOLCF ***
  3338 
  3339 * Theory Representable defines a class "rep" of domains that are
  3340 representable (via an ep-pair) in the universal domain type "udom".
  3341 Instances are provided for all type constructors defined in HOLCF.
  3342 
  3343 * The 'new_domain' command is a purely definitional version of the
  3344 domain package, for representable domains.  Syntax is identical to the
  3345 old domain package.  The 'new_domain' package also supports indirect
  3346 recursion using previously-defined type constructors.  See
  3347 src/HOLCF/ex/New_Domain.thy for examples.
  3348 
  3349 * Method "fixrec_simp" unfolds one step of a fixrec-defined constant
  3350 on the left-hand side of an equation, and then performs
  3351 simplification.  Rewriting is done using rules declared with the
  3352 "fixrec_simp" attribute.  The "fixrec_simp" method is intended as a
  3353 replacement for "fixpat"; see src/HOLCF/ex/Fixrec_ex.thy for examples.
  3354 
  3355 * The pattern-match compiler in 'fixrec' can now handle constructors
  3356 with HOL function types.  Pattern-match combinators for the Pair
  3357 constructor are pre-configured.
  3358 
  3359 * The 'fixrec' package now produces better fixed-point induction rules
  3360 for mutually-recursive definitions:  Induction rules have conclusions
  3361 of the form "P foo bar" instead of "P <foo, bar>".
  3362 
  3363 * The constant "sq_le" (with infix syntax "<<" or "\<sqsubseteq>") has
  3364 been renamed to "below".  The name "below" now replaces "less" in many
  3365 theorem names.  (Legacy theorem names using "less" are still supported
  3366 as well.)
  3367 
  3368 * The 'fixrec' package now supports "bottom patterns".  Bottom
  3369 patterns can be used to generate strictness rules, or to make
  3370 functions more strict (much like the bang-patterns supported by the
  3371 Glasgow Haskell Compiler).  See src/HOLCF/ex/Fixrec_ex.thy for
  3372 examples.
  3373 
  3374 
  3375 *** ML ***
  3376 
  3377 * Support for Poly/ML 5.3.0, with improved reporting of compiler
  3378 errors and run-time exceptions, including detailed source positions.
  3379 
  3380 * Structure Name_Space (formerly NameSpace) now manages uniquely
  3381 identified entries, with some additional information such as source
  3382 position, logical grouping etc.
  3383 
  3384 * Theory and context data is now introduced by the simplified and
  3385 modernized functors Theory_Data, Proof_Data, Generic_Data.  Data needs
  3386 to be pure, but the old TheoryDataFun for mutable data (with explicit
  3387 copy operation) is still available for some time.
  3388 
  3389 * Structure Synchronized (cf. src/Pure/Concurrent/synchronized.ML)
  3390 provides a high-level programming interface to synchronized state
  3391 variables with atomic update.  This works via pure function
  3392 application within a critical section -- its runtime should be as
  3393 short as possible; beware of deadlocks if critical code is nested,
  3394 either directly or indirectly via other synchronized variables!
  3395 
  3396 * Structure Unsynchronized (cf. src/Pure/ML-Systems/unsynchronized.ML)
  3397 wraps raw ML references, explicitly indicating their non-thread-safe
  3398 behaviour.  The Isar toplevel keeps this structure open, to
  3399 accommodate Proof General as well as quick and dirty interactive
  3400 experiments with references.
  3401 
  3402 * PARALLEL_CHOICE and PARALLEL_GOALS provide basic support for
  3403 parallel tactical reasoning.
  3404 
  3405 * Tacticals Subgoal.FOCUS, Subgoal.FOCUS_PREMS, Subgoal.FOCUS_PARAMS
  3406 are similar to SUBPROOF, but are slightly more flexible: only the
  3407 specified parts of the subgoal are imported into the context, and the
  3408 body tactic may introduce new subgoals and schematic variables.
  3409 
  3410 * Old tactical METAHYPS, which does not observe the proof context, has
  3411 been renamed to Old_Goals.METAHYPS and awaits deletion.  Use SUBPROOF
  3412 or Subgoal.FOCUS etc.
  3413 
  3414 * Renamed functor TableFun to Table, and GraphFun to Graph.  (Since
  3415 functors have their own ML name space there is no point to mark them
  3416 separately.)  Minor INCOMPATIBILITY.
  3417 
  3418 * Renamed NamedThmsFun to Named_Thms.  INCOMPATIBILITY.
  3419 
  3420 * Renamed several structures FooBar to Foo_Bar.  Occasional,
  3421 INCOMPATIBILITY.
  3422 
  3423 * Operations of structure Skip_Proof no longer require quick_and_dirty
  3424 mode, which avoids critical setmp.
  3425 
  3426 * Eliminated old Attrib.add_attributes, Method.add_methods and related
  3427 combinators for "args".  INCOMPATIBILITY, need to use simplified
  3428 Attrib/Method.setup introduced in Isabelle2009.
  3429 
  3430 * Proper context for simpset_of, claset_of, clasimpset_of.  May fall
  3431 back on global_simpset_of, global_claset_of, global_clasimpset_of as
  3432 last resort.  INCOMPATIBILITY.
  3433 
  3434 * Display.pretty_thm now requires a proper context (cf. former
  3435 ProofContext.pretty_thm).  May fall back on Display.pretty_thm_global
  3436 or even Display.pretty_thm_without_context as last resort.
  3437 INCOMPATIBILITY.
  3438 
  3439 * Discontinued Display.pretty_ctyp/cterm etc.  INCOMPATIBILITY, use
  3440 Syntax.pretty_typ/term directly, preferably with proper context
  3441 instead of global theory.
  3442 
  3443 
  3444 *** System ***
  3445 
  3446 * Further fine tuning of parallel proof checking, scales up to 8 cores
  3447 (max. speedup factor 5.0).  See also Goal.parallel_proofs in ML and
  3448 usedir option -q.
  3449 
  3450 * Support for additional "Isabelle components" via etc/components, see
  3451 also the system manual.
  3452 
  3453 * The isabelle makeall tool now operates on all components with
  3454 IsaMakefile, not just hardwired "logics".
  3455 
  3456 * Removed "compress" option from isabelle-process and isabelle usedir;
  3457 this is always enabled.
  3458 
  3459 * Discontinued support for Poly/ML 4.x versions.
  3460 
  3461 * Isabelle tool "wwwfind" provides web interface for 'find_theorems'
  3462 on a given logic image.  This requires the lighttpd webserver and is
  3463 currently supported on Linux only.
  3464 
  3465 
  3466 
  3467 New in Isabelle2009 (April 2009)
  3468 --------------------------------
  3469 
  3470 *** General ***
  3471 
  3472 * Simplified main Isabelle executables, with less surprises on
  3473 case-insensitive file-systems (such as Mac OS).
  3474 
  3475   - The main Isabelle tool wrapper is now called "isabelle" instead of
  3476     "isatool."
  3477 
  3478   - The former "isabelle" alias for "isabelle-process" has been
  3479     removed (should rarely occur to regular users).
  3480 
  3481   - The former "isabelle-interface" and its alias "Isabelle" have been
  3482     removed (interfaces are now regular Isabelle tools).
  3483 
  3484 Within scripts and make files, the Isabelle environment variables
  3485 ISABELLE_TOOL and ISABELLE_PROCESS replace old ISATOOL and ISABELLE,
  3486 respectively.  (The latter are still available as legacy feature.)
  3487 
  3488 The old isabelle-interface wrapper could react in confusing ways if
  3489 the interface was uninstalled or changed otherwise.  Individual
  3490 interface tool configuration is now more explicit, see also the
  3491 Isabelle system manual.  In particular, Proof General is now available
  3492 via "isabelle emacs".
  3493 
  3494 INCOMPATIBILITY, need to adapt derivative scripts.  Users may need to
  3495 purge installed copies of Isabelle executables and re-run "isabelle
  3496 install -p ...", or use symlinks.
  3497 
  3498 * The default for ISABELLE_HOME_USER is now ~/.isabelle instead of the
  3499 old ~/isabelle, which was slightly non-standard and apt to cause
  3500 surprises on case-insensitive file-systems (such as Mac OS).
  3501 
  3502 INCOMPATIBILITY, need to move existing ~/isabelle/etc,
  3503 ~/isabelle/heaps, ~/isabelle/browser_info to the new place.  Special
  3504 care is required when using older releases of Isabelle.  Note that
  3505 ISABELLE_HOME_USER can be changed in Isabelle/etc/settings of any
  3506 Isabelle distribution, in order to use the new ~/.isabelle uniformly.
  3507 
  3508 * Proofs of fully specified statements are run in parallel on
  3509 multi-core systems.  A speedup factor of 2.5 to 3.2 can be expected on
  3510 a regular 4-core machine, if the initial heap space is made reasonably
  3511 large (cf. Poly/ML option -H).  (Requires Poly/ML 5.2.1 or later.)
  3512 
  3513 * The main reference manuals ("isar-ref", "implementation", and
  3514 "system") have been updated and extended.  Formally checked references
  3515 as hyperlinks are now available uniformly.
  3516 
  3517 
  3518 *** Pure ***
  3519 
  3520 * Complete re-implementation of locales.  INCOMPATIBILITY in several
  3521 respects.  The most important changes are listed below.  See the
  3522 Tutorial on Locales ("locales" manual) for details.
  3523 
  3524 - In locale expressions, instantiation replaces renaming.  Parameters
  3525 must be declared in a for clause.  To aid compatibility with previous
  3526 parameter inheritance, in locale declarations, parameters that are not
  3527 'touched' (instantiation position "_" or omitted) are implicitly added
  3528 with their syntax at the beginning of the for clause.
  3529 
  3530 - Syntax from abbreviations and definitions in locales is available in
  3531 locale expressions and context elements.  The latter is particularly
  3532 useful in locale declarations.
  3533 
  3534 - More flexible mechanisms to qualify names generated by locale
  3535 expressions.  Qualifiers (prefixes) may be specified in locale
  3536 expressions, and can be marked as mandatory (syntax: "name!:") or
  3537 optional (syntax "name?:").  The default depends for plain "name:"
  3538 depends on the situation where a locale expression is used: in
  3539 commands 'locale' and 'sublocale' prefixes are optional, in
  3540 'interpretation' and 'interpret' prefixes are mandatory.  The old
  3541 implicit qualifiers derived from the parameter names of a locale are
  3542 no longer generated.
  3543 
  3544 - Command "sublocale l < e" replaces "interpretation l < e".  The
  3545 instantiation clause in "interpretation" and "interpret" (square
  3546 brackets) is no longer available.  Use locale expressions.
  3547 
  3548 - When converting proof scripts, mandatory qualifiers in
  3549 'interpretation' and 'interpret' should be retained by default, even
  3550 if this is an INCOMPATIBILITY compared to former behavior.  In the
  3551 worst case, use the "name?:" form for non-mandatory ones.  Qualifiers
  3552 in locale expressions range over a single locale instance only.
  3553 
  3554 - Dropped locale element "includes".  This is a major INCOMPATIBILITY.
  3555 In existing theorem specifications replace the includes element by the
  3556 respective context elements of the included locale, omitting those
  3557 that are already present in the theorem specification.  Multiple
  3558 assume elements of a locale should be replaced by a single one
  3559 involving the locale predicate.  In the proof body, declarations (most
  3560 notably theorems) may be regained by interpreting the respective
  3561 locales in the proof context as required (command "interpret").
  3562 
  3563 If using "includes" in replacement of a target solely because the
  3564 parameter types in the theorem are not as general as in the target,
  3565 consider declaring a new locale with additional type constraints on
  3566 the parameters (context element "constrains").
  3567 
  3568 - Discontinued "locale (open)".  INCOMPATIBILITY.
  3569 
  3570 - Locale interpretation commands no longer attempt to simplify goal.
  3571 INCOMPATIBILITY: in rare situations the generated goal differs.  Use
  3572 methods intro_locales and unfold_locales to clarify.
  3573 
  3574 - Locale interpretation commands no longer accept interpretation
  3575 attributes.  INCOMPATIBILITY.
  3576 
  3577 * Class declaration: so-called "base sort" must not be given in import
  3578 list any longer, but is inferred from the specification.  Particularly
  3579 in HOL, write
  3580 
  3581     class foo = ...
  3582 
  3583 instead of
  3584 
  3585     class foo = type + ...
  3586 
  3587 * Class target: global versions of theorems stemming do not carry a
  3588 parameter prefix any longer.  INCOMPATIBILITY.
  3589 
  3590 * Class 'instance' command no longer accepts attached definitions.
  3591 INCOMPATIBILITY, use proper 'instantiation' target instead.
  3592 
  3593 * Recovered hiding of consts, which was accidentally broken in
  3594 Isabelle2007.  Potential INCOMPATIBILITY, ``hide const c'' really
  3595 makes c inaccessible; consider using ``hide (open) const c'' instead.
  3596 
  3597 * Slightly more coherent Pure syntax, with updated documentation in
  3598 isar-ref manual.  Removed locales meta_term_syntax and
  3599 meta_conjunction_syntax: TERM and &&& (formerly &&) are now permanent,
  3600 INCOMPATIBILITY in rare situations.  Note that &&& should not be used
  3601 directly in regular applications.
  3602 
  3603 * There is a new syntactic category "float_const" for signed decimal
  3604 fractions (e.g. 123.45 or -123.45).
  3605 
  3606 * Removed exotic 'token_translation' command.  INCOMPATIBILITY, use ML
  3607 interface with 'setup' command instead.
  3608 
  3609 * Command 'local_setup' is similar to 'setup', but operates on a local
  3610 theory context.
  3611 
  3612 * The 'axiomatization' command now only works within a global theory
  3613 context.  INCOMPATIBILITY.
  3614 
  3615 * Goal-directed proof now enforces strict proof irrelevance wrt. sort
  3616 hypotheses.  Sorts required in the course of reasoning need to be
  3617 covered by the constraints in the initial statement, completed by the
  3618 type instance information of the background theory.  Non-trivial sort
  3619 hypotheses, which rarely occur in practice, may be specified via
  3620 vacuous propositions of the form SORT_CONSTRAINT('a::c).  For example:
  3621 
  3622   lemma assumes "SORT_CONSTRAINT('a::empty)" shows False ...
  3623 
  3624 The result contains an implicit sort hypotheses as before --
  3625 SORT_CONSTRAINT premises are eliminated as part of the canonical rule
  3626 normalization.
  3627 
  3628 * Generalized Isar history, with support for linear undo, direct state
  3629 addressing etc.
  3630 
  3631 * Changed defaults for unify configuration options:
  3632 
  3633   unify_trace_bound = 50 (formerly 25)
  3634   unify_search_bound = 60 (formerly 30)
  3635 
  3636 * Different bookkeeping for code equations (INCOMPATIBILITY):
  3637 
  3638   a) On theory merge, the last set of code equations for a particular
  3639      constant is taken (in accordance with the policy applied by other
  3640      parts of the code generator framework).
  3641 
  3642   b) Code equations stemming from explicit declarations (e.g. code
  3643      attribute) gain priority over default code equations stemming
  3644      from definition, primrec, fun etc.
  3645 
  3646 * Keyword 'code_exception' now named 'code_abort'.  INCOMPATIBILITY.
  3647 
  3648 * Unified theorem tables for both code generators.  Thus [code
  3649 func] has disappeared and only [code] remains.  INCOMPATIBILITY.
  3650 
  3651 * Command 'find_consts' searches for constants based on type and name
  3652 patterns, e.g.
  3653 
  3654     find_consts "_ => bool"
  3655 
  3656 By default, matching is against subtypes, but it may be restricted to
  3657 the whole type.  Searching by name is possible.  Multiple queries are
  3658 conjunctive and queries may be negated by prefixing them with a
  3659 hyphen:
  3660 
  3661     find_consts strict: "_ => bool" name: "Int" -"int => int"
  3662 
  3663 * New 'find_theorems' criterion "solves" matches theorems that
  3664 directly solve the current goal (modulo higher-order unification).
  3665 
  3666 * Auto solve feature for main theorem statements: whenever a new goal
  3667 is stated, "find_theorems solves" is called; any theorems that could
  3668 solve the lemma directly are listed as part of the goal state.
  3669 Cf. associated options in Proof General Isabelle settings menu,
  3670 enabled by default, with reasonable timeout for pathological cases of
  3671 higher-order unification.
  3672 
  3673 
  3674 *** Document preparation ***
  3675 
  3676 * Antiquotation @{lemma} now imitates a regular terminal proof,
  3677 demanding keyword 'by' and supporting the full method expression
  3678 syntax just like the Isar command 'by'.
  3679 
  3680 
  3681 *** HOL ***
  3682 
  3683 * Integrated main parts of former image HOL-Complex with HOL.  Entry
  3684 points Main and Complex_Main remain as before.
  3685 
  3686 * Logic image HOL-Plain provides a minimal HOL with the most important
  3687 tools available (inductive, datatype, primrec, ...).  This facilitates
  3688 experimentation and tool development.  Note that user applications
  3689 (and library theories) should never refer to anything below theory
  3690 Main, as before.
  3691 
  3692 * Logic image HOL-Main stops at theory Main, and thus facilitates
  3693 experimentation due to shorter build times.
  3694 
  3695 * Logic image HOL-NSA contains theories of nonstandard analysis which
  3696 were previously part of former HOL-Complex.  Entry point Hyperreal
  3697 remains valid, but theories formerly using Complex_Main should now use
  3698 new entry point Hypercomplex.
  3699 
  3700 * Generic ATP manager for Sledgehammer, based on ML threads instead of
  3701 Posix processes.  Avoids potentially expensive forking of the ML
  3702 process.  New thread-based implementation also works on non-Unix
  3703 platforms (Cygwin).  Provers are no longer hardwired, but defined
  3704 within the theory via plain ML wrapper functions.  Basic Sledgehammer
  3705 commands are covered in the isar-ref manual.
  3706 
  3707 * Wrapper scripts for remote SystemOnTPTP service allows to use
  3708 sledgehammer without local ATP installation (Vampire etc.). Other
  3709 provers may be included via suitable ML wrappers, see also
  3710 src/HOL/ATP_Linkup.thy.
  3711 
  3712 * ATP selection (E/Vampire/Spass) is now via Proof General's settings
  3713 menu.
  3714 
  3715 * The metis method no longer fails because the theorem is too trivial
  3716 (contains the empty clause).
  3717 
  3718 * The metis method now fails in the usual manner, rather than raising
  3719 an exception, if it determines that it cannot prove the theorem.
  3720 
  3721 * Method "coherent" implements a prover for coherent logic (see also
  3722 src/Tools/coherent.ML).
  3723 
  3724 * Constants "undefined" and "default" replace "arbitrary".  Usually
  3725 "undefined" is the right choice to replace "arbitrary", though
  3726 logically there is no difference.  INCOMPATIBILITY.
  3727 
  3728 * Command "value" now integrates different evaluation mechanisms.  The
  3729 result of the first successful evaluation mechanism is printed.  In
  3730 square brackets a particular named evaluation mechanisms may be
  3731 specified (currently, [SML], [code] or [nbe]).  See further
  3732 src/HOL/ex/Eval_Examples.thy.
  3733 
  3734 * Normalization by evaluation now allows non-leftlinear equations.
  3735 Declare with attribute [code nbe].
  3736 
  3737 * Methods "case_tac" and "induct_tac" now refer to the very same rules
  3738 as the structured Isar versions "cases" and "induct", cf. the
  3739 corresponding "cases" and "induct" attributes.  Mutual induction rules
  3740 are now presented as a list of individual projections
  3741 (e.g. foo_bar.inducts for types foo and bar); the old format with
  3742 explicit HOL conjunction is no longer supported.  INCOMPATIBILITY, in
  3743 rare situations a different rule is selected --- notably nested tuple
  3744 elimination instead of former prod.exhaust: use explicit (case_tac t
  3745 rule: prod.exhaust) here.
  3746 
  3747 * Attributes "cases", "induct", "coinduct" support "del" option.
  3748 
  3749 * Removed fact "case_split_thm", which duplicates "case_split".
  3750 
  3751 * The option datatype has been moved to a new theory Option.  Renamed
  3752 option_map to Option.map, and o2s to Option.set, INCOMPATIBILITY.
  3753 
  3754 * New predicate "strict_mono" classifies strict functions on partial
  3755 orders.  With strict functions on linear orders, reasoning about
  3756 (in)equalities is facilitated by theorems "strict_mono_eq",
  3757 "strict_mono_less_eq" and "strict_mono_less".
  3758 
  3759 * Some set operations are now proper qualified constants with
  3760 authentic syntax.  INCOMPATIBILITY:
  3761 
  3762     op Int ~>   Set.Int
  3763     op Un ~>    Set.Un
  3764     INTER ~>    Set.INTER
  3765     UNION ~>    Set.UNION
  3766     Inter ~>    Set.Inter
  3767     Union ~>    Set.Union
  3768     {} ~>       Set.empty
  3769     UNIV ~>     Set.UNIV
  3770 
  3771 * Class complete_lattice with operations Inf, Sup, INFI, SUPR now in
  3772 theory Set.
  3773 
  3774 * Auxiliary class "itself" has disappeared -- classes without any
  3775 parameter are treated as expected by the 'class' command.
  3776 
  3777 * Leibnitz's Series for Pi and the arcus tangens and logarithm series.
  3778 
  3779 * Common decision procedures (Cooper, MIR, Ferrack, Approximation,
  3780 Dense_Linear_Order) are now in directory HOL/Decision_Procs.
  3781 
  3782 * Theory src/HOL/Decision_Procs/Approximation provides the new proof
  3783 method "approximation".  It proves formulas on real values by using
  3784 interval arithmetic.  In the formulas are also the transcendental
  3785 functions sin, cos, tan, atan, ln, exp and the constant pi are
  3786 allowed. For examples see
  3787 src/HOL/Descision_Procs/ex/Approximation_Ex.thy.
  3788 
  3789 * Theory "Reflection" now resides in HOL/Library.
  3790 
  3791 * Entry point to Word library now simply named "Word".
  3792 INCOMPATIBILITY.
  3793 
  3794 * Made source layout more coherent with logical distribution
  3795 structure:
  3796 
  3797     src/HOL/Library/RType.thy ~> src/HOL/Typerep.thy
  3798     src/HOL/Library/Code_Message.thy ~> src/HOL/
  3799     src/HOL/Library/GCD.thy ~> src/HOL/
  3800     src/HOL/Library/Order_Relation.thy ~> src/HOL/
  3801     src/HOL/Library/Parity.thy ~> src/HOL/
  3802     src/HOL/Library/Univ_Poly.thy ~> src/HOL/
  3803     src/HOL/Real/ContNotDenum.thy ~> src/HOL/Library/
  3804     src/HOL/Real/Lubs.thy ~> src/HOL/
  3805     src/HOL/Real/PReal.thy ~> src/HOL/
  3806     src/HOL/Real/Rational.thy ~> src/HOL/
  3807     src/HOL/Real/RComplete.thy ~> src/HOL/
  3808     src/HOL/Real/RealDef.thy ~> src/HOL/
  3809     src/HOL/Real/RealPow.thy ~> src/HOL/
  3810     src/HOL/Real/Real.thy ~> src/HOL/
  3811     src/HOL/Complex/Complex_Main.thy ~> src/HOL/
  3812     src/HOL/Complex/Complex.thy ~> src/HOL/
  3813     src/HOL/Complex/FrechetDeriv.thy ~> src/HOL/Library/
  3814     src/HOL/Complex/Fundamental_Theorem_Algebra.thy ~> src/HOL/Library/
  3815     src/HOL/Hyperreal/Deriv.thy ~> src/HOL/
  3816     src/HOL/Hyperreal/Fact.thy ~> src/HOL/
  3817     src/HOL/Hyperreal/Integration.thy ~> src/HOL/
  3818     src/HOL/Hyperreal/Lim.thy ~> src/HOL/
  3819     src/HOL/Hyperreal/Ln.thy ~> src/HOL/
  3820     src/HOL/Hyperreal/Log.thy ~> src/HOL/
  3821     src/HOL/Hyperreal/MacLaurin.thy ~> src/HOL/
  3822     src/HOL/Hyperreal/NthRoot.thy ~> src/HOL/
  3823     src/HOL/Hyperreal/Series.thy ~> src/HOL/
  3824     src/HOL/Hyperreal/SEQ.thy ~> src/HOL/
  3825     src/HOL/Hyperreal/Taylor.thy ~> src/HOL/
  3826     src/HOL/Hyperreal/Transcendental.thy ~> src/HOL/
  3827     src/HOL/Real/Float ~> src/HOL/Library/
  3828     src/HOL/Real/HahnBanach ~> src/HOL/HahnBanach
  3829     src/HOL/Real/RealVector.thy ~> src/HOL/
  3830 
  3831     src/HOL/arith_data.ML ~> src/HOL/Tools
  3832     src/HOL/hologic.ML ~> src/HOL/Tools
  3833     src/HOL/simpdata.ML ~> src/HOL/Tools
  3834     src/HOL/int_arith1.ML ~> src/HOL/Tools/int_arith.ML
  3835     src/HOL/int_factor_simprocs.ML ~> src/HOL/Tools
  3836     src/HOL/nat_simprocs.ML ~> src/HOL/Tools
  3837     src/HOL/Real/float_arith.ML ~> src/HOL/Tools
  3838     src/HOL/Real/float_syntax.ML ~> src/HOL/Tools
  3839     src/HOL/Real/rat_arith.ML ~> src/HOL/Tools
  3840     src/HOL/Real/real_arith.ML ~> src/HOL/Tools
  3841 
  3842     src/HOL/Library/Array.thy ~> src/HOL/Imperative_HOL
  3843     src/HOL/Library/Heap_Monad.thy ~> src/HOL/Imperative_HOL
  3844     src/HOL/Library/Heap.thy ~> src/HOL/Imperative_HOL
  3845     src/HOL/Library/Imperative_HOL.thy ~> src/HOL/Imperative_HOL
  3846     src/HOL/Library/Ref.thy ~> src/HOL/Imperative_HOL
  3847     src/HOL/Library/Relational.thy ~> src/HOL/Imperative_HOL
  3848 
  3849 * If methods "eval" and "evaluation" encounter a structured proof
  3850 state with !!/==>, only the conclusion is evaluated to True (if
  3851 possible), avoiding strange error messages.
  3852 
  3853 * Method "sizechange" automates termination proofs using (a
  3854 modification of) the size-change principle.  Requires SAT solver.  See
  3855 src/HOL/ex/Termination.thy for examples.
  3856 
  3857 * Simplifier: simproc for let expressions now unfolds if bound
  3858 variable occurs at most once in let expression body.  INCOMPATIBILITY.
  3859 
  3860 * Method "arith": Linear arithmetic now ignores all inequalities when
  3861 fast_arith_neq_limit is exceeded, instead of giving up entirely.
  3862 
  3863 * New attribute "arith" for facts that should always be used
  3864 automatically by arithmetic. It is intended to be used locally in
  3865 proofs, e.g.
  3866 
  3867   assumes [arith]: "x > 0"
  3868 
  3869 Global usage is discouraged because of possible performance impact.
  3870 
  3871 * New classes "top" and "bot" with corresponding operations "top" and
  3872 "bot" in theory Orderings; instantiation of class "complete_lattice"
  3873 requires instantiation of classes "top" and "bot".  INCOMPATIBILITY.
  3874 
  3875 * Changed definition lemma "less_fun_def" in order to provide an
  3876 instance for preorders on functions; use lemma "less_le" instead.
  3877 INCOMPATIBILITY.
  3878 
  3879 * Theory Orderings: class "wellorder" moved here, with explicit
  3880 induction rule "less_induct" as assumption.  For instantiation of
  3881 "wellorder" by means of predicate "wf", use rule wf_wellorderI.
  3882 INCOMPATIBILITY.
  3883 
  3884 * Theory Orderings: added class "preorder" as superclass of "order".
  3885 INCOMPATIBILITY: Instantiation proofs for order, linorder
  3886 etc. slightly changed.  Some theorems named order_class.* now named
  3887 preorder_class.*.
  3888 
  3889 * Theory Relation: renamed "refl" to "refl_on", "reflexive" to "refl,
  3890 "diag" to "Id_on".
  3891 
  3892 * Theory Finite_Set: added a new fold combinator of type
  3893 
  3894   ('a => 'b => 'b) => 'b => 'a set => 'b
  3895 
  3896 Occasionally this is more convenient than the old fold combinator
  3897 which is now defined in terms of the new one and renamed to
  3898 fold_image.
  3899 
  3900 * Theories Ring_and_Field and OrderedGroup: The lemmas "group_simps"
  3901 and "ring_simps" have been replaced by "algebra_simps" (which can be
  3902 extended with further lemmas!).  At the moment both still exist but
  3903 the former will disappear at some point.
  3904 
  3905 * Theory Power: Lemma power_Suc is now declared as a simp rule in
  3906 class recpower.  Type-specific simp rules for various recpower types
  3907 have been removed.  INCOMPATIBILITY, rename old lemmas as follows:
  3908 
  3909 rat_power_0    -> power_0
  3910 rat_power_Suc  -> power_Suc
  3911 realpow_0      -> power_0
  3912 realpow_Suc    -> power_Suc
  3913 complexpow_0   -> power_0
  3914 complexpow_Suc -> power_Suc
  3915 power_poly_0   -> power_0
  3916 power_poly_Suc -> power_Suc
  3917 
  3918 * Theories Ring_and_Field and Divides: Definition of "op dvd" has been
  3919 moved to separate class dvd in Ring_and_Field; a couple of lemmas on
  3920 dvd has been generalized to class comm_semiring_1.  Likewise a bunch
  3921 of lemmas from Divides has been generalized from nat to class
  3922 semiring_div.  INCOMPATIBILITY.  This involves the following theorem
  3923 renames resulting from duplicate elimination:
  3924 
  3925     dvd_def_mod ~>          dvd_eq_mod_eq_0
  3926     zero_dvd_iff ~>         dvd_0_left_iff
  3927     dvd_0 ~>                dvd_0_right
  3928     DIVISION_BY_ZERO_DIV ~> div_by_0
  3929     DIVISION_BY_ZERO_MOD ~> mod_by_0
  3930     mult_div ~>             div_mult_self2_is_id
  3931     mult_mod ~>             mod_mult_self2_is_0
  3932 
  3933 * Theory IntDiv: removed many lemmas that are instances of class-based
  3934 generalizations (from Divides and Ring_and_Field).  INCOMPATIBILITY,
  3935 rename old lemmas as follows:
  3936 
  3937 dvd_diff               -> nat_dvd_diff
  3938 dvd_zminus_iff         -> dvd_minus_iff
  3939 mod_add1_eq            -> mod_add_eq
  3940 mod_mult1_eq           -> mod_mult_right_eq
  3941 mod_mult1_eq'          -> mod_mult_left_eq
  3942 mod_mult_distrib_mod   -> mod_mult_eq
  3943 nat_mod_add_left_eq    -> mod_add_left_eq
  3944 nat_mod_add_right_eq   -> mod_add_right_eq
  3945 nat_mod_div_trivial    -> mod_div_trivial
  3946 nat_mod_mod_trivial    -> mod_mod_trivial
  3947 zdiv_zadd_self1        -> div_add_self1
  3948 zdiv_zadd_self2        -> div_add_self2
  3949 zdiv_zmult_self1       -> div_mult_self2_is_id
  3950 zdiv_zmult_self2       -> div_mult_self1_is_id
  3951 zdvd_triv_left         -> dvd_triv_left
  3952 zdvd_triv_right        -> dvd_triv_right
  3953 zdvd_zmult_cancel_disj -> dvd_mult_cancel_left
  3954 zmod_eq0_zdvd_iff      -> dvd_eq_mod_eq_0[symmetric]
  3955 zmod_zadd_left_eq      -> mod_add_left_eq
  3956 zmod_zadd_right_eq     -> mod_add_right_eq
  3957 zmod_zadd_self1        -> mod_add_self1
  3958 zmod_zadd_self2        -> mod_add_self2
  3959 zmod_zadd1_eq          -> mod_add_eq
  3960 zmod_zdiff1_eq         -> mod_diff_eq
  3961 zmod_zdvd_zmod         -> mod_mod_cancel
  3962 zmod_zmod_cancel       -> mod_mod_cancel
  3963 zmod_zmult_self1       -> mod_mult_self2_is_0
  3964 zmod_zmult_self2       -> mod_mult_self1_is_0
  3965 zmod_1                 -> mod_by_1
  3966 zdiv_1                 -> div_by_1
  3967 zdvd_abs1              -> abs_dvd_iff
  3968 zdvd_abs2              -> dvd_abs_iff
  3969 zdvd_refl              -> dvd_refl
  3970 zdvd_trans             -> dvd_trans
  3971 zdvd_zadd              -> dvd_add
  3972 zdvd_zdiff             -> dvd_diff
  3973 zdvd_zminus_iff        -> dvd_minus_iff
  3974 zdvd_zminus2_iff       -> minus_dvd_iff
  3975 zdvd_zmultD            -> dvd_mult_right
  3976 zdvd_zmultD2           -> dvd_mult_left
  3977 zdvd_zmult_mono        -> mult_dvd_mono
  3978 zdvd_0_right           -> dvd_0_right
  3979 zdvd_0_left            -> dvd_0_left_iff
  3980 zdvd_1_left            -> one_dvd
  3981 zminus_dvd_iff         -> minus_dvd_iff
  3982 
  3983 * Theory Rational: 'Fract k 0' now equals '0'.  INCOMPATIBILITY.
  3984 
  3985 * The real numbers offer decimal input syntax: 12.34 is translated
  3986 into 1234/10^2. This translation is not reversed upon output.
  3987 
  3988 * Theory Library/Polynomial defines an abstract type 'a poly of
  3989 univariate polynomials with coefficients of type 'a.  In addition to
  3990 the standard ring operations, it also supports div and mod.  Code
  3991 generation is also supported, using list-style constructors.
  3992 
  3993 * Theory Library/Inner_Product defines a class of real_inner for real
  3994 inner product spaces, with an overloaded operation inner :: 'a => 'a
  3995 => real.  Class real_inner is a subclass of real_normed_vector from
  3996 theory RealVector.
  3997 
  3998 * Theory Library/Product_Vector provides instances for the product
  3999 type 'a * 'b of several classes from RealVector and Inner_Product.
  4000 Definitions of addition, subtraction, scalar multiplication, norms,
  4001 and inner products are included.
  4002 
  4003 * Theory Library/Bit defines the field "bit" of integers modulo 2.  In
  4004 addition to the field operations, numerals and case syntax are also
  4005 supported.
  4006 
  4007 * Theory Library/Diagonalize provides constructive version of Cantor's
  4008 first diagonalization argument.
  4009 
  4010 * Theory Library/GCD: Curried operations gcd, lcm (for nat) and zgcd,
  4011 zlcm (for int); carried together from various gcd/lcm developements in
  4012 the HOL Distribution.  Constants zgcd and zlcm replace former igcd and
  4013 ilcm; corresponding theorems renamed accordingly.  INCOMPATIBILITY,
  4014 may recover tupled syntax as follows:
  4015 
  4016     hide (open) const gcd
  4017     abbreviation gcd where
  4018       "gcd == (%(a, b). GCD.gcd a b)"
  4019     notation (output)
  4020       GCD.gcd ("gcd '(_, _')")
  4021 
  4022 The same works for lcm, zgcd, zlcm.
  4023 
  4024 * Theory Library/Nat_Infinity: added addition, numeral syntax and more
  4025 instantiations for algebraic structures.  Removed some duplicate
  4026 theorems.  Changes in simp rules.  INCOMPATIBILITY.
  4027 
  4028 * ML antiquotation @{code} takes a constant as argument and generates
  4029 corresponding code in background and inserts name of the corresponding
  4030 resulting ML value/function/datatype constructor binding in place.
  4031 All occurrences of @{code} with a single ML block are generated
  4032 simultaneously.  Provides a generic and safe interface for
  4033 instrumentalizing code generation.  See
  4034 src/HOL/Decision_Procs/Ferrack.thy for a more ambitious application.
  4035 In future you ought to refrain from ad-hoc compiling generated SML
  4036 code on the ML toplevel.  Note that (for technical reasons) @{code}
  4037 cannot refer to constants for which user-defined serializations are
  4038 set.  Refer to the corresponding ML counterpart directly in that
  4039 cases.
  4040 
  4041 * Command 'rep_datatype': instead of theorem names the command now
  4042 takes a list of terms denoting the constructors of the type to be
  4043 represented as datatype.  The characteristic theorems have to be
  4044 proven.  INCOMPATIBILITY.  Also observe that the following theorems
  4045 have disappeared in favour of existing ones:
  4046 
  4047     unit_induct                 ~> unit.induct
  4048     prod_induct                 ~> prod.induct
  4049     sum_induct                  ~> sum.induct
  4050     Suc_Suc_eq                  ~> nat.inject
  4051     Suc_not_Zero Zero_not_Suc   ~> nat.distinct
  4052 
  4053 
  4054 *** HOL-Algebra ***
  4055 
  4056 * New locales for orders and lattices where the equivalence relation
  4057 is not restricted to equality.  INCOMPATIBILITY: all order and lattice
  4058 locales use a record structure with field eq for the equivalence.
  4059 
  4060 * New theory of factorial domains.
  4061 
  4062 * Units_l_inv and Units_r_inv are now simp rules by default.
  4063 INCOMPATIBILITY.  Simplifier proof that require deletion of l_inv
  4064 and/or r_inv will now also require deletion of these lemmas.
  4065 
  4066 * Renamed the following theorems, INCOMPATIBILITY:
  4067 
  4068 UpperD ~> Upper_memD
  4069 LowerD ~> Lower_memD
  4070 least_carrier ~> least_closed
  4071 greatest_carrier ~> greatest_closed
  4072 greatest_Lower_above ~> greatest_Lower_below
  4073 one_zero ~> carrier_one_zero
  4074 one_not_zero ~> carrier_one_not_zero  (collision with assumption)
  4075 
  4076 
  4077 *** HOL-Nominal ***
  4078 
  4079 * Nominal datatypes can now contain type-variables.
  4080 
  4081 * Commands 'nominal_inductive' and 'equivariance' work with local
  4082 theory targets.
  4083 
  4084 * Nominal primrec can now works with local theory targets and its
  4085 specification syntax now conforms to the general format as seen in
  4086 'inductive' etc.
  4087 
  4088 * Method "perm_simp" honours the standard simplifier attributes
  4089 (no_asm), (no_asm_use) etc.
  4090 
  4091 * The new predicate #* is defined like freshness, except that on the
  4092 left hand side can be a set or list of atoms.
  4093 
  4094 * Experimental command 'nominal_inductive2' derives strong induction
  4095 principles for inductive definitions.  In contrast to
  4096 'nominal_inductive', which can only deal with a fixed number of
  4097 binders, it can deal with arbitrary expressions standing for sets of
  4098 atoms to be avoided.  The only inductive definition we have at the
  4099 moment that needs this generalisation is the typing rule for Lets in
  4100 the algorithm W:
  4101 
  4102  Gamma |- t1 : T1   (x,close Gamma T1)::Gamma |- t2 : T2   x#Gamma
  4103  -----------------------------------------------------------------
  4104          Gamma |- Let x be t1 in t2 : T2
  4105 
  4106 In this rule one wants to avoid all the binders that are introduced by
  4107 "close Gamma T1".  We are looking for other examples where this
  4108 feature might be useful.  Please let us know.
  4109 
  4110 
  4111 *** HOLCF ***
  4112 
  4113 * Reimplemented the simplification procedure for proving continuity
  4114 subgoals.  The new simproc is extensible; users can declare additional
  4115 continuity introduction rules with the attribute [cont2cont].
  4116 
  4117 * The continuity simproc now uses a different introduction rule for
  4118 solving continuity subgoals on terms with lambda abstractions.  In
  4119 some rare cases the new simproc may fail to solve subgoals that the
  4120 old one could solve, and "simp add: cont2cont_LAM" may be necessary.
  4121 Potential INCOMPATIBILITY.
  4122 
  4123 * Command 'fixrec': specification syntax now conforms to the general
  4124 format as seen in 'inductive' etc.  See src/HOLCF/ex/Fixrec_ex.thy for
  4125 examples.  INCOMPATIBILITY.
  4126 
  4127 
  4128 *** ZF ***
  4129 
  4130 * Proof of Zorn's Lemma for partial orders.
  4131 
  4132 
  4133 *** ML ***
  4134 
  4135 * Multithreading for Poly/ML 5.1/5.2 is no longer supported, only for
  4136 Poly/ML 5.2.1 or later.  Important note: the TimeLimit facility
  4137 depends on multithreading, so timouts will not work before Poly/ML
  4138 5.2.1!
  4139 
  4140 * High-level support for concurrent ML programming, see
  4141 src/Pure/Cuncurrent.  The data-oriented model of "future values" is
  4142 particularly convenient to organize independent functional
  4143 computations.  The concept of "synchronized variables" provides a
  4144 higher-order interface for components with shared state, avoiding the
  4145 delicate details of mutexes and condition variables.  (Requires
  4146 Poly/ML 5.2.1 or later.)
  4147 
  4148 * ML bindings produced via Isar commands are stored within the Isar
  4149 context (theory or proof).  Consequently, commands like 'use' and 'ML'
  4150 become thread-safe and work with undo as expected (concerning
  4151 top-level bindings, not side-effects on global references).
  4152 INCOMPATIBILITY, need to provide proper Isar context when invoking the
  4153 compiler at runtime; really global bindings need to be given outside a
  4154 theory.  (Requires Poly/ML 5.2 or later.)
  4155 
  4156 * Command 'ML_prf' is analogous to 'ML' but works within a proof
  4157 context.  Top-level ML bindings are stored within the proof context in
  4158 a purely sequential fashion, disregarding the nested proof structure.
  4159 ML bindings introduced by 'ML_prf' are discarded at the end of the
  4160 proof.  (Requires Poly/ML 5.2 or later.)
  4161 
  4162 * Simplified ML attribute and method setup, cf. functions Attrib.setup
  4163 and Method.setup, as well as Isar commands 'attribute_setup' and
  4164 'method_setup'.  INCOMPATIBILITY for 'method_setup', need to simplify
  4165 existing code accordingly, or use plain 'setup' together with old
  4166 Method.add_method.
  4167 
  4168 * Simplified ML oracle interface Thm.add_oracle promotes 'a -> cterm
  4169 to 'a -> thm, while results are always tagged with an authentic oracle
  4170 name.  The Isar command 'oracle' is now polymorphic, no argument type
  4171 is specified.  INCOMPATIBILITY, need to simplify existing oracle code
  4172 accordingly.  Note that extra performance may be gained by producing
  4173 the cterm carefully, avoiding slow Thm.cterm_of.
  4174 
  4175 * Simplified interface for defining document antiquotations via
  4176 ThyOutput.antiquotation, ThyOutput.output, and optionally
  4177 ThyOutput.maybe_pretty_source.  INCOMPATIBILITY, need to simplify user
  4178 antiquotations accordingly, see src/Pure/Thy/thy_output.ML for common
  4179 examples.
  4180 
  4181 * More systematic treatment of long names, abstract name bindings, and
  4182 name space operations.  Basic operations on qualified names have been
  4183 move from structure NameSpace to Long_Name, e.g. Long_Name.base_name,
  4184 Long_Name.append.  Old type bstring has been mostly replaced by
  4185 abstract type binding (see structure Binding), which supports precise
  4186 qualification by packages and local theory targets, as well as proper
  4187 tracking of source positions.  INCOMPATIBILITY, need to wrap old
  4188 bstring values into Binding.name, or better pass through abstract
  4189 bindings everywhere.  See further src/Pure/General/long_name.ML,
  4190 src/Pure/General/binding.ML and src/Pure/General/name_space.ML
  4191 
  4192 * Result facts (from PureThy.note_thms, ProofContext.note_thms,
  4193 LocalTheory.note etc.) now refer to the *full* internal name, not the
  4194 bstring as before.  INCOMPATIBILITY, not detected by ML type-checking!
  4195 
  4196 * Disposed old type and term read functions (Sign.read_def_typ,
  4197 Sign.read_typ, Sign.read_def_terms, Sign.read_term,
  4198 Thm.read_def_cterms, Thm.read_cterm etc.).  INCOMPATIBILITY, should
  4199 use regular Syntax.read_typ, Syntax.read_term, Syntax.read_typ_global,
  4200 Syntax.read_term_global etc.; see also OldGoals.read_term as last
  4201 resort for legacy applications.
  4202 
  4203 * Disposed old declarations, tactics, tactic combinators that refer to
  4204 the simpset or claset of an implicit theory (such as Addsimps,
  4205 Simp_tac, SIMPSET).  INCOMPATIBILITY, should use @{simpset} etc. in
  4206 embedded ML text, or local_simpset_of with a proper context passed as
  4207 explicit runtime argument.
  4208 
  4209 * Rules and tactics that read instantiations (read_instantiate,
  4210 res_inst_tac, thin_tac, subgoal_tac etc.) now demand a proper proof
  4211 context, which is required for parsing and type-checking.  Moreover,
  4212 the variables are specified as plain indexnames, not string encodings
  4213 thereof.  INCOMPATIBILITY.
  4214 
  4215 * Generic Toplevel.add_hook interface allows to analyze the result of
  4216 transactions.  E.g. see src/Pure/ProofGeneral/proof_general_pgip.ML
  4217 for theorem dependency output of transactions resulting in a new
  4218 theory state.
  4219 
  4220 * ML antiquotations: block-structured compilation context indicated by
  4221 \<lbrace> ... \<rbrace>; additional antiquotation forms:
  4222 
  4223   @{binding name}                         - basic name binding
  4224   @{let ?pat = term}                      - term abbreviation (HO matching)
  4225   @{note name = fact}                     - fact abbreviation
  4226   @{thm fact}                             - singleton fact (with attributes)
  4227   @{thms fact}                            - general fact (with attributes)
  4228   @{lemma prop by method}                 - singleton goal
  4229   @{lemma prop by meth1 meth2}            - singleton goal
  4230   @{lemma prop1 ... propN by method}      - general goal
  4231   @{lemma prop1 ... propN by meth1 meth2} - general goal
  4232   @{lemma (open) ...}                     - open derivation
  4233 
  4234 
  4235 *** System ***
  4236 
  4237 * The Isabelle "emacs" tool provides a specific interface to invoke
  4238 Proof General / Emacs, with more explicit failure if that is not
  4239 installed (the old isabelle-interface script silently falls back on
  4240 isabelle-process).  The PROOFGENERAL_HOME setting determines the
  4241 installation location of the Proof General distribution.
  4242 
  4243 * Isabelle/lib/classes/Pure.jar provides basic support to integrate
  4244 the Isabelle process into a JVM/Scala application.  See
  4245 Isabelle/lib/jedit/plugin for a minimal example.  (The obsolete Java
  4246 process wrapper has been discontinued.)
  4247 
  4248 * Added homegrown Isabelle font with unicode layout, see lib/fonts.
  4249 
  4250 * Various status messages (with exact source position information) are
  4251 emitted, if proper markup print mode is enabled.  This allows
  4252 user-interface components to provide detailed feedback on internal
  4253 prover operations.
  4254 
  4255 
  4256 
  4257 New in Isabelle2008 (June 2008)
  4258 -------------------------------
  4259 
  4260 *** General ***
  4261 
  4262 * The Isabelle/Isar Reference Manual (isar-ref) has been reorganized
  4263 and updated, with formally checked references as hyperlinks.
  4264 
  4265 * Theory loader: use_thy (and similar operations) no longer set the
  4266 implicit ML context, which was occasionally hard to predict and in
  4267 conflict with concurrency.  INCOMPATIBILITY, use ML within Isar which
  4268 provides a proper context already.
  4269 
  4270 * Theory loader: old-style ML proof scripts being *attached* to a thy
  4271 file are no longer supported.  INCOMPATIBILITY, regular 'uses' and
  4272 'use' within a theory file will do the job.
  4273 
  4274 * Name space merge now observes canonical order, i.e. the second space
  4275 is inserted into the first one, while existing entries in the first
  4276 space take precedence.  INCOMPATIBILITY in rare situations, may try to
  4277 swap theory imports.
  4278 
  4279 * Syntax: symbol \<chi> is now considered a letter.  Potential
  4280 INCOMPATIBILITY in identifier syntax etc.
  4281 
  4282 * Outer syntax: string tokens no longer admit escaped white space,
  4283 which was an accidental (undocumented) feature.  INCOMPATIBILITY, use
  4284 white space without escapes.
  4285 
  4286 * Outer syntax: string tokens may contain arbitrary character codes
  4287 specified via 3 decimal digits (as in SML).  E.g. "foo\095bar" for
  4288 "foo_bar".
  4289 
  4290 
  4291 *** Pure ***
  4292 
  4293 * Context-dependent token translations.  Default setup reverts locally
  4294 fixed variables, and adds hilite markup for undeclared frees.
  4295 
  4296 * Unused theorems can be found using the new command 'unused_thms'.
  4297 There are three ways of invoking it:
  4298 
  4299 (1) unused_thms
  4300      Only finds unused theorems in the current theory.
  4301 
  4302 (2) unused_thms thy_1 ... thy_n -
  4303      Finds unused theorems in the current theory and all of its ancestors,
  4304      excluding the theories thy_1 ... thy_n and all of their ancestors.
  4305 
  4306 (3) unused_thms thy_1 ... thy_n - thy'_1 ... thy'_m
  4307      Finds unused theorems in the theories thy'_1 ... thy'_m and all of
  4308      their ancestors, excluding the theories thy_1 ... thy_n and all of
  4309      their ancestors.
  4310 
  4311 In order to increase the readability of the list produced by
  4312 unused_thms, theorems that have been created by a particular instance
  4313 of a theory command such as 'inductive' or 'function' are considered
  4314 to belong to the same "group", meaning that if at least one theorem in
  4315 this group is used, the other theorems in the same group are no longer
  4316 reported as unused.  Moreover, if all theorems in the group are
  4317 unused, only one theorem in the group is displayed.
  4318 
  4319 Note that proof objects have to be switched on in order for
  4320 unused_thms to work properly (i.e. !proofs must be >= 1, which is
  4321 usually the case when using Proof General with the default settings).
  4322 
  4323 * Authentic naming of facts disallows ad-hoc overwriting of previous
  4324 theorems within the same name space.  INCOMPATIBILITY, need to remove
  4325 duplicate fact bindings, or even accidental fact duplications.  Note
  4326 that tools may maintain dynamically scoped facts systematically, using
  4327 PureThy.add_thms_dynamic.
  4328 
  4329 * Command 'hide' now allows to hide from "fact" name space as well.
  4330 
  4331 * Eliminated destructive theorem database, simpset, claset, and
  4332 clasimpset.  Potential INCOMPATIBILITY, really need to observe linear
  4333 update of theories within ML code.
  4334 
  4335 * Eliminated theory ProtoPure and CPure, leaving just one Pure theory.
  4336 INCOMPATIBILITY, object-logics depending on former Pure require
  4337 additional setup PureThy.old_appl_syntax_setup; object-logics
  4338 depending on former CPure need to refer to Pure.
  4339 
  4340 * Commands 'use' and 'ML' are now purely functional, operating on
  4341 theory/local_theory.  Removed former 'ML_setup' (on theory), use 'ML'
  4342 instead.  Added 'ML_val' as mere diagnostic replacement for 'ML'.
  4343 INCOMPATIBILITY.
  4344 
  4345 * Command 'setup': discontinued implicit version with ML reference.
  4346 
  4347 * Instantiation target allows for simultaneous specification of class
  4348 instance operations together with an instantiation proof.
  4349 Type-checking phase allows to refer to class operations uniformly.
  4350 See src/HOL/Complex/Complex.thy for an Isar example and
  4351 src/HOL/Library/Eval.thy for an ML example.
  4352 
  4353 * Indexing of literal facts: be more serious about including only
  4354 facts from the visible specification/proof context, but not the
  4355 background context (locale etc.).  Affects `prop` notation and method
  4356 "fact".  INCOMPATIBILITY: need to name facts explicitly in rare
  4357 situations.
  4358 
  4359 * Method "cases", "induct", "coinduct": removed obsolete/undocumented
  4360 "(open)" option, which used to expose internal bound variables to the
  4361 proof text.
  4362 
  4363 * Isar statements: removed obsolete case "rule_context".
  4364 INCOMPATIBILITY, better use explicit fixes/assumes.
  4365 
  4366 * Locale proofs: default proof step now includes 'unfold_locales';
  4367 hence 'proof' without argument may be used to unfold locale
  4368 predicates.
  4369 
  4370 
  4371 *** Document preparation ***
  4372 
  4373 * Simplified pdfsetup.sty: color/hyperref is used unconditionally for
  4374 both pdf and dvi (hyperlinks usually work in xdvi as well); removed
  4375 obsolete thumbpdf setup (contemporary PDF viewers do this on the
  4376 spot); renamed link color from "darkblue" to "linkcolor" (default
  4377 value unchanged, can be redefined via \definecolor); no longer sets
  4378 "a4paper" option (unnecessary or even intrusive).
  4379 
  4380 * Antiquotation @{lemma A method} proves proposition A by the given
  4381 method (either a method name or a method name plus (optional) method
  4382 arguments in parentheses) and prints A just like @{prop A}.
  4383 
  4384 
  4385 *** HOL ***
  4386 
  4387 * New primrec package.  Specification syntax conforms in style to
  4388 definition/function/....  No separate induction rule is provided.  The
  4389 "primrec" command distinguishes old-style and new-style specifications
  4390 by syntax.  The former primrec package is now named OldPrimrecPackage.
  4391 When adjusting theories, beware: constants stemming from new-style
  4392 primrec specifications have authentic syntax.
  4393 
  4394 * Metis prover is now an order of magnitude faster, and also works
  4395 with multithreading.
  4396 
  4397 * Metis: the maximum number of clauses that can be produced from a
  4398 theorem is now given by the attribute max_clauses.  Theorems that
  4399 exceed this number are ignored, with a warning printed.
  4400 
  4401 * Sledgehammer no longer produces structured proofs by default. To
  4402 enable, declare [[sledgehammer_full = true]].  Attributes
  4403 reconstruction_modulus, reconstruction_sorts renamed
  4404 sledgehammer_modulus, sledgehammer_sorts.  INCOMPATIBILITY.
  4405 
  4406 * Method "induct_scheme" derives user-specified induction rules
  4407 from well-founded induction and completeness of patterns. This factors
  4408 out some operations that are done internally by the function package
  4409 and makes them available separately.  See
  4410 src/HOL/ex/Induction_Scheme.thy for examples.
  4411 
  4412 * More flexible generation of measure functions for termination
  4413 proofs: Measure functions can be declared by proving a rule of the
  4414 form "is_measure f" and giving it the [measure_function] attribute.
  4415 The "is_measure" predicate is logically meaningless (always true), and
  4416 just guides the heuristic.  To find suitable measure functions, the
  4417 termination prover sets up the goal "is_measure ?f" of the appropriate
  4418 type and generates all solutions by prolog-style backwards proof using
  4419 the declared rules.
  4420 
  4421 This setup also deals with rules like 
  4422 
  4423   "is_measure f ==> is_measure (list_size f)"
  4424 
  4425 which accommodates nested datatypes that recurse through lists.
  4426 Similar rules are predeclared for products and option types.
  4427 
  4428 * Turned the type of sets "'a set" into an abbreviation for "'a => bool"
  4429 
  4430   INCOMPATIBILITIES:
  4431 
  4432   - Definitions of overloaded constants on sets have to be replaced by
  4433     definitions on => and bool.
  4434 
  4435   - Some definitions of overloaded operators on sets can now be proved
  4436     using the definitions of the operators on => and bool.  Therefore,
  4437     the following theorems have been renamed:
  4438 
  4439       subset_def   -> subset_eq
  4440       psubset_def  -> psubset_eq
  4441       set_diff_def -> set_diff_eq
  4442       Compl_def    -> Compl_eq
  4443       Sup_set_def  -> Sup_set_eq
  4444       Inf_set_def  -> Inf_set_eq
  4445       sup_set_def  -> sup_set_eq
  4446       inf_set_def  -> inf_set_eq
  4447 
  4448   - Due to the incompleteness of the HO unification algorithm, some
  4449     rules such as subst may require manual instantiation, if some of
  4450     the unknowns in the rule is a set.
  4451 
  4452   - Higher order unification and forward proofs:
  4453     The proof pattern
  4454 
  4455       have "P (S::'a set)" <...>
  4456       then have "EX S. P S" ..
  4457 
  4458     no longer works (due to the incompleteness of the HO unification
  4459     algorithm) and must be replaced by the pattern
  4460 
  4461       have "EX S. P S"
  4462       proof
  4463         show "P S" <...>
  4464       qed
  4465 
  4466   - Calculational reasoning with subst (or similar rules):
  4467     The proof pattern
  4468 
  4469       have "P (S::'a set)" <...>
  4470       also have "S = T" <...>
  4471       finally have "P T" .
  4472 
  4473     no longer works (for similar reasons as the previous example) and
  4474     must be replaced by something like
  4475 
  4476       have "P (S::'a set)" <...>
  4477       moreover have "S = T" <...>
  4478       ultimately have "P T" by simp
  4479 
  4480   - Tactics or packages written in ML code:
  4481     Code performing pattern matching on types via
  4482 
  4483       Type ("set", [T]) => ...
  4484 
  4485     must be rewritten. Moreover, functions like strip_type or
  4486     binder_types no longer return the right value when applied to a
  4487     type of the form
  4488 
  4489       T1 => ... => Tn => U => bool
  4490 
  4491     rather than
  4492 
  4493       T1 => ... => Tn => U set
  4494 
  4495 * Merged theories Wellfounded_Recursion, Accessible_Part and
  4496 Wellfounded_Relations to theory Wellfounded.
  4497 
  4498 * Explicit class "eq" for executable equality.  INCOMPATIBILITY.
  4499 
  4500 * Class finite no longer treats UNIV as class parameter.  Use class
  4501 enum from theory Library/Enum instead to achieve a similar effect.
  4502 INCOMPATIBILITY.
  4503 
  4504 * Theory List: rule list_induct2 now has explicitly named cases "Nil"
  4505 and "Cons".  INCOMPATIBILITY.
  4506 
  4507 * HOL (and FOL): renamed variables in rules imp_elim and swap.
  4508 Potential INCOMPATIBILITY.
  4509 
  4510 * Theory Product_Type: duplicated lemmas split_Pair_apply and
  4511 injective_fst_snd removed, use split_eta and prod_eqI instead.
  4512 Renamed upd_fst to apfst and upd_snd to apsnd.  INCOMPATIBILITY.
  4513 
  4514 * Theory Nat: removed redundant lemmas that merely duplicate lemmas of
  4515 the same name in theory Orderings:
  4516 
  4517   less_trans
  4518   less_linear
  4519   le_imp_less_or_eq
  4520   le_less_trans
  4521   less_le_trans
  4522   less_not_sym
  4523   less_asym
  4524 
  4525 Renamed less_imp_le to less_imp_le_nat, and less_irrefl to
  4526 less_irrefl_nat.  Potential INCOMPATIBILITY due to more general types
  4527 and different variable names.
  4528 
  4529 * Library/Option_ord.thy: Canonical order on option type.
  4530 
  4531 * Library/RBT.thy: Red-black trees, an efficient implementation of
  4532 finite maps.
  4533 
  4534 * Library/Countable.thy: Type class for countable types.
  4535 
  4536 * Theory Int: The representation of numerals has changed.  The infix
  4537 operator BIT and the bit datatype with constructors B0 and B1 have
  4538 disappeared.  INCOMPATIBILITY, use "Int.Bit0 x" and "Int.Bit1 y" in
  4539 place of "x BIT bit.B0" and "y BIT bit.B1", respectively.  Theorems
  4540 involving BIT, B0, or B1 have been renamed with "Bit0" or "Bit1"
  4541 accordingly.
  4542 
  4543 * Theory Nat: definition of <= and < on natural numbers no longer
  4544 depend on well-founded relations.  INCOMPATIBILITY.  Definitions
  4545 le_def and less_def have disappeared.  Consider lemmas not_less
  4546 [symmetric, where ?'a = nat] and less_eq [symmetric] instead.
  4547 
  4548 * Theory Finite_Set: locales ACf, ACe, ACIf, ACIfSL and ACIfSLlin
  4549 (whose purpose mainly is for various fold_set functionals) have been
  4550 abandoned in favor of the existing algebraic classes
  4551 ab_semigroup_mult, comm_monoid_mult, ab_semigroup_idem_mult,
  4552 lower_semilattice (resp. upper_semilattice) and linorder.
  4553 INCOMPATIBILITY.
  4554 
  4555 * Theory Transitive_Closure: induct and cases rules now declare proper
  4556 case_names ("base" and "step").  INCOMPATIBILITY.
  4557 
  4558 * Theorem Inductive.lfp_ordinal_induct generalized to complete
  4559 lattices.  The form set-specific version is available as
  4560 Inductive.lfp_ordinal_induct_set.
  4561 
  4562 * Renamed theorems "power.simps" to "power_int.simps".
  4563 INCOMPATIBILITY.
  4564 
  4565 * Class semiring_div provides basic abstract properties of semirings
  4566 with division and modulo operations.  Subsumes former class dvd_mod.
  4567 
  4568 * Merged theories IntDef, Numeral and IntArith into unified theory
  4569 Int.  INCOMPATIBILITY.
  4570 
  4571 * Theory Library/Code_Index: type "index" now represents natural
  4572 numbers rather than integers.  INCOMPATIBILITY.
  4573 
  4574 * New class "uminus" with operation "uminus" (split of from class
  4575 "minus" which now only has operation "minus", binary).
  4576 INCOMPATIBILITY.
  4577 
  4578 * Constants "card", "internal_split", "option_map" now with authentic
  4579 syntax.  INCOMPATIBILITY.
  4580 
  4581 * Definitions subset_def, psubset_def, set_diff_def, Compl_def,
  4582 le_bool_def, less_bool_def, le_fun_def, less_fun_def, inf_bool_def,
  4583 sup_bool_def, Inf_bool_def, Sup_bool_def, inf_fun_def, sup_fun_def,
  4584 Inf_fun_def, Sup_fun_def, inf_set_def, sup_set_def, Inf_set_def,
  4585 Sup_set_def, le_def, less_def, option_map_def now with object
  4586 equality.  INCOMPATIBILITY.
  4587 
  4588 * Records. Removed K_record, and replaced it by pure lambda term
  4589 %x. c. The simplifier setup is now more robust against eta expansion.
  4590 INCOMPATIBILITY: in cases explicitly referring to K_record.
  4591 
  4592 * Library/Multiset: {#a, b, c#} abbreviates {#a#} + {#b#} + {#c#}.
  4593 
  4594 * Library/ListVector: new theory of arithmetic vector operations.
  4595 
  4596 * Library/Order_Relation: new theory of various orderings as sets of
  4597 pairs.  Defines preorders, partial orders, linear orders and
  4598 well-orders on sets and on types.
  4599 
  4600 
  4601 *** ZF ***
  4602 
  4603 * Renamed some theories to allow to loading both ZF and HOL in the
  4604 same session:
  4605 
  4606   Datatype  -> Datatype_ZF
  4607   Inductive -> Inductive_ZF
  4608   Int       -> Int_ZF
  4609   IntDiv    -> IntDiv_ZF
  4610   Nat       -> Nat_ZF
  4611   List      -> List_ZF
  4612   Main      -> Main_ZF
  4613 
  4614 INCOMPATIBILITY: ZF theories that import individual theories below
  4615 Main might need to be adapted.  Regular theory Main is still
  4616 available, as trivial extension of Main_ZF.
  4617 
  4618 
  4619 *** ML ***
  4620 
  4621 * ML within Isar: antiquotation @{const name} or @{const
  4622 name(typargs)} produces statically-checked Const term.
  4623 
  4624 * Functor NamedThmsFun: data is available to the user as dynamic fact
  4625 (of the same name).  Removed obsolete print command.
  4626 
  4627 * Removed obsolete "use_legacy_bindings" function.
  4628 
  4629 * The ``print mode'' is now a thread-local value derived from a global
  4630 template (the former print_mode reference), thus access becomes
  4631 non-critical.  The global print_mode reference is for session
  4632 management only; user-code should use print_mode_value,
  4633 print_mode_active, PrintMode.setmp etc.  INCOMPATIBILITY.
  4634 
  4635 * Functions system/system_out provide a robust way to invoke external
  4636 shell commands, with propagation of interrupts (requires Poly/ML
  4637 5.2.1).  Do not use OS.Process.system etc. from the basis library!
  4638 
  4639 
  4640 *** System ***
  4641 
  4642 * Default settings: PROOFGENERAL_OPTIONS no longer impose xemacs ---
  4643 in accordance with Proof General 3.7, which prefers GNU emacs.
  4644 
  4645 * isatool tty runs Isabelle process with plain tty interaction;
  4646 optional line editor may be specified via ISABELLE_LINE_EDITOR
  4647 setting, the default settings attempt to locate "ledit" and "rlwrap".
  4648 
  4649 * isatool browser now works with Cygwin as well, using general
  4650 "javapath" function defined in Isabelle process environment.
  4651 
  4652 * YXML notation provides a simple and efficient alternative to
  4653 standard XML transfer syntax.  See src/Pure/General/yxml.ML and
  4654 isatool yxml as described in the Isabelle system manual.
  4655 
  4656 * JVM class isabelle.IsabelleProcess (located in Isabelle/lib/classes)
  4657 provides general wrapper for managing an Isabelle process in a robust
  4658 fashion, with ``cooked'' output from stdin/stderr.
  4659 
  4660 * Rudimentary Isabelle plugin for jEdit (see Isabelle/lib/jedit),
  4661 based on Isabelle/JVM process wrapper (see Isabelle/lib/classes).
  4662 
  4663 * Removed obsolete THIS_IS_ISABELLE_BUILD feature.  NB: the documented
  4664 way of changing the user's settings is via
  4665 ISABELLE_HOME_USER/etc/settings, which is a fully featured bash
  4666 script.
  4667 
  4668 * Multithreading.max_threads := 0 refers to the number of actual CPU
  4669 cores of the underlying machine, which is a good starting point for
  4670 optimal performance tuning.  The corresponding usedir option -M allows
  4671 "max" as an alias for "0".  WARNING: does not work on certain versions
  4672 of Mac OS (with Poly/ML 5.1).
  4673 
  4674 * isabelle-process: non-ML sessions are run with "nice", to reduce the
  4675 adverse effect of Isabelle flooding interactive front-ends (notably
  4676 ProofGeneral / XEmacs).
  4677 
  4678 
  4679 
  4680 New in Isabelle2007 (November 2007)
  4681 -----------------------------------
  4682 
  4683 *** General ***
  4684 
  4685 * More uniform information about legacy features, notably a
  4686 warning/error of "Legacy feature: ...", depending on the state of the
  4687 tolerate_legacy_features flag (default true). FUTURE INCOMPATIBILITY:
  4688 legacy features will disappear eventually.
  4689 
  4690 * Theory syntax: the header format ``theory A = B + C:'' has been
  4691 discontinued in favour of ``theory A imports B C begin''.  Use isatool
  4692 fixheaders to convert existing theory files.  INCOMPATIBILITY.
  4693 
  4694 * Theory syntax: the old non-Isar theory file format has been
  4695 discontinued altogether.  Note that ML proof scripts may still be used
  4696 with Isar theories; migration is usually quite simple with the ML
  4697 function use_legacy_bindings.  INCOMPATIBILITY.
  4698 
  4699 * Theory syntax: some popular names (e.g. 'class', 'declaration',
  4700 'fun', 'help', 'if') are now keywords.  INCOMPATIBILITY, use double
  4701 quotes.
  4702 
  4703 * Theory loader: be more serious about observing the static theory
  4704 header specifications (including optional directories), but not the
  4705 accidental file locations of previously successful loads.  The strict
  4706 update policy of former update_thy is now already performed by
  4707 use_thy, so the former has been removed; use_thys updates several
  4708 theories simultaneously, just as 'imports' within a theory header
  4709 specification, but without merging the results.  Potential
  4710 INCOMPATIBILITY: may need to refine theory headers and commands
  4711 ROOT.ML which depend on load order.
  4712 
  4713 * Theory loader: optional support for content-based file
  4714 identification, instead of the traditional scheme of full physical
  4715 path plus date stamp; configured by the ISABELLE_FILE_IDENT setting
  4716 (cf. the system manual).  The new scheme allows to work with
  4717 non-finished theories in persistent session images, such that source
  4718 files may be moved later on without requiring reloads.
  4719 
  4720 * Theory loader: old-style ML proof scripts being *attached* to a thy
  4721 file (with the same base name as the theory) are considered a legacy
  4722 feature, which will disappear eventually. Even now, the theory loader
  4723 no longer maintains dependencies on such files.
  4724 
  4725 * Syntax: the scope for resolving ambiguities via type-inference is
  4726 now limited to individual terms, instead of whole simultaneous
  4727 specifications as before. This greatly reduces the complexity of the
  4728 syntax module and improves flexibility by separating parsing and
  4729 type-checking. INCOMPATIBILITY: additional type-constraints (explicit
  4730 'fixes' etc.) are required in rare situations.
  4731 
  4732 * Syntax: constants introduced by new-style packages ('definition',
  4733 'abbreviation' etc.) are passed through the syntax module in
  4734 ``authentic mode''. This means that associated mixfix annotations
  4735 really stick to such constants, independently of potential name space
  4736 ambiguities introduced later on. INCOMPATIBILITY: constants in parse
  4737 trees are represented slightly differently, may need to adapt syntax
  4738 translations accordingly. Use CONST marker in 'translations' and
  4739 @{const_syntax} antiquotation in 'parse_translation' etc.
  4740 
  4741 * Legacy goal package: reduced interface to the bare minimum required
  4742 to keep existing proof scripts running.  Most other user-level
  4743 functions are now part of the OldGoals structure, which is *not* open
  4744 by default (consider isatool expandshort before open OldGoals).
  4745 Removed top_sg, prin, printyp, pprint_term/typ altogether, because
  4746 these tend to cause confusion about the actual goal (!) context being
  4747 used here, which is not necessarily the same as the_context().
  4748 
  4749 * Command 'find_theorems': supports "*" wild-card in "name:"
  4750 criterion; "with_dups" option.  Certain ProofGeneral versions might
  4751 support a specific search form (see ProofGeneral/CHANGES).
  4752 
  4753 * The ``prems limit'' option (cf. ProofContext.prems_limit) is now -1
  4754 by default, which means that "prems" (and also "fixed variables") are
  4755 suppressed from proof state output.  Note that the ProofGeneral
  4756 settings mechanism allows to change and save options persistently, but
  4757 older versions of Isabelle will fail to start up if a negative prems
  4758 limit is imposed.
  4759 
  4760 * Local theory targets may be specified by non-nested blocks of
  4761 ``context/locale/class ... begin'' followed by ``end''.  The body may
  4762 contain definitions, theorems etc., including any derived mechanism
  4763 that has been implemented on top of these primitives.  This concept
  4764 generalizes the existing ``theorem (in ...)'' towards more versatility
  4765 and scalability.
  4766 
  4767 * Proof General interface: proper undo of final 'end' command;
  4768 discontinued Isabelle/classic mode (ML proof scripts).
  4769 
  4770 
  4771 *** Document preparation ***
  4772 
  4773 * Added antiquotation @{theory name} which prints the given name,
  4774 after checking that it refers to a valid ancestor theory in the
  4775 current context.
  4776 
  4777 * Added antiquotations @{ML_type text} and @{ML_struct text} which
  4778 check the given source text as ML type/structure, printing verbatim.
  4779 
  4780 * Added antiquotation @{abbrev "c args"} which prints the abbreviation
  4781 "c args == rhs" given in the current context.  (Any number of
  4782 arguments may be given on the LHS.)
  4783 
  4784 
  4785 *** Pure ***
  4786 
  4787 * The 'class' package offers a combination of axclass and locale to
  4788 achieve Haskell-like type classes in Isabelle.  Definitions and
  4789 theorems within a class context produce both relative results (with
  4790 implicit parameters according to the locale context), and polymorphic
  4791 constants with qualified polymorphism (according to the class
  4792 context).  Within the body context of a 'class' target, a separate
  4793 syntax layer ("user space type system") takes care of converting
  4794 between global polymorphic consts and internal locale representation.
  4795 See src/HOL/ex/Classpackage.thy for examples (as well as main HOL).
  4796 "isatool doc classes" provides a tutorial.
  4797 
  4798 * Generic code generator framework allows to generate executable
  4799 code for ML and Haskell (including Isabelle classes).  A short usage
  4800 sketch:
  4801 
  4802     internal compilation:
  4803         export_code <list of constants (term syntax)> in SML
  4804     writing SML code to a file:
  4805         export_code <list of constants (term syntax)> in SML <filename>
  4806     writing OCaml code to a file:
  4807         export_code <list of constants (term syntax)> in OCaml <filename>
  4808     writing Haskell code to a bunch of files:
  4809         export_code <list of constants (term syntax)> in Haskell <filename>
  4810 
  4811     evaluating closed propositions to True/False using code generation:
  4812         method ``eval''
  4813 
  4814 Reasonable default setup of framework in HOL.
  4815 
  4816 Theorem attributs for selecting and transforming function equations theorems:
  4817 
  4818     [code fun]:        select a theorem as function equation for a specific constant
  4819     [code fun del]:    deselect a theorem as function equation for a specific constant
  4820     [code inline]:     select an equation theorem for unfolding (inlining) in place
  4821     [code inline del]: deselect an equation theorem for unfolding (inlining) in place
  4822 
  4823 User-defined serializations (target in {SML, OCaml, Haskell}):
  4824 
  4825     code_const <and-list of constants (term syntax)>
  4826       {(target) <and-list of const target syntax>}+
  4827 
  4828     code_type <and-list of type constructors>
  4829       {(target) <and-list of type target syntax>}+
  4830 
  4831     code_instance <and-list of instances>
  4832       {(target)}+
  4833         where instance ::= <type constructor> :: <class>
  4834 
  4835     code_class <and_list of classes>
  4836       {(target) <and-list of class target syntax>}+
  4837         where class target syntax ::= <class name> {where {<classop> == <target syntax>}+}?
  4838 
  4839 code_instance and code_class only are effective to target Haskell.
  4840 
  4841 For example usage see src/HOL/ex/Codegenerator.thy and
  4842 src/HOL/ex/Codegenerator_Pretty.thy.  A separate tutorial on code
  4843 generation from Isabelle/HOL theories is available via "isatool doc
  4844 codegen".
  4845 
  4846 * Code generator: consts in 'consts_code' Isar commands are now
  4847 referred to by usual term syntax (including optional type
  4848 annotations).
  4849 
  4850 * Command 'no_translations' removes translation rules from theory
  4851 syntax.
  4852 
  4853 * Overloaded definitions are now actually checked for acyclic
  4854 dependencies.  The overloading scheme is slightly more general than
  4855 that of Haskell98, although Isabelle does not demand an exact
  4856 correspondence to type class and instance declarations.
  4857 INCOMPATIBILITY, use ``defs (unchecked overloaded)'' to admit more
  4858 exotic versions of overloading -- at the discretion of the user!
  4859 
  4860 Polymorphic constants are represented via type arguments, i.e. the
  4861 instantiation that matches an instance against the most general
  4862 declaration given in the signature.  For example, with the declaration
  4863 c :: 'a => 'a => 'a, an instance c :: nat => nat => nat is represented
  4864 as c(nat).  Overloading is essentially simultaneous structural
  4865 recursion over such type arguments.  Incomplete specification patterns
  4866 impose global constraints on all occurrences, e.g. c('a * 'a) on the
  4867 LHS means that more general c('a * 'b) will be disallowed on any RHS.
  4868 Command 'print_theory' outputs the normalized system of recursive
  4869 equations, see section "definitions".
  4870 
  4871 * Configuration options are maintained within the theory or proof
  4872 context (with name and type bool/int/string), providing a very simple
  4873 interface to a poor-man's version of general context data.  Tools may
  4874 declare options in ML (e.g. using Attrib.config_int) and then refer to
  4875 these values using Config.get etc.  Users may change options via an
  4876 associated attribute of the same name.  This form of context
  4877 declaration works particularly well with commands 'declare' or
  4878 'using', for example ``declare [[foo = 42]]''.  Thus it has become
  4879 very easy to avoid global references, which would not observe Isar
  4880 toplevel undo/redo and fail to work with multithreading.
  4881 
  4882 Various global ML references of Pure and HOL have been turned into
  4883 configuration options:
  4884 
  4885   Unify.search_bound		unify_search_bound
  4886   Unify.trace_bound		unify_trace_bound
  4887   Unify.trace_simp		unify_trace_simp
  4888   Unify.trace_types		unify_trace_types
  4889   Simplifier.simp_depth_limit	simp_depth_limit
  4890   Blast.depth_limit		blast_depth_limit
  4891   DatatypeProp.dtK		datatype_distinctness_limit
  4892   fast_arith_neq_limit  	fast_arith_neq_limit
  4893   fast_arith_split_limit	fast_arith_split_limit
  4894 
  4895 * Named collections of theorems may be easily installed as context
  4896 data using the functor NamedThmsFun (see also
  4897 src/Pure/Tools/named_thms.ML).  The user may add or delete facts via
  4898 attributes; there is also a toplevel print command.  This facility is
  4899 just a common case of general context data, which is the preferred way
  4900 for anything more complex than just a list of facts in canonical
  4901 order.
  4902 
  4903 * Isar: command 'declaration' augments a local theory by generic
  4904 declaration functions written in ML.  This enables arbitrary content
  4905 being added to the context, depending on a morphism that tells the
  4906 difference of the original declaration context wrt. the application
  4907 context encountered later on.
  4908 
  4909 * Isar: proper interfaces for simplification procedures.  Command
  4910 'simproc_setup' declares named simprocs (with match patterns, and body
  4911 text in ML).  Attribute "simproc" adds/deletes simprocs in the current
  4912 context.  ML antiquotation @{simproc name} retrieves named simprocs.
  4913 
  4914 * Isar: an extra pair of brackets around attribute declarations
  4915 abbreviates a theorem reference involving an internal dummy fact,
  4916 which will be ignored later --- only the effect of the attribute on
  4917 the background context will persist.  This form of in-place
  4918 declarations is particularly useful with commands like 'declare' and
  4919 'using', for example ``have A using [[simproc a]] by simp''.
  4920 
  4921 * Isar: method "assumption" (and implicit closing of subproofs) now
  4922 takes simple non-atomic goal assumptions into account: after applying
  4923 an assumption as a rule the resulting subgoals are solved by atomic
  4924 assumption steps.  This is particularly useful to finish 'obtain'
  4925 goals, such as "!!x. (!!x. P x ==> thesis) ==> P x ==> thesis",
  4926 without referring to the original premise "!!x. P x ==> thesis" in the
  4927 Isar proof context.  POTENTIAL INCOMPATIBILITY: method "assumption" is
  4928 more permissive.
  4929 
  4930 * Isar: implicit use of prems from the Isar proof context is
  4931 considered a legacy feature.  Common applications like ``have A .''
  4932 may be replaced by ``have A by fact'' or ``note `A`''.  In general,
  4933 referencing facts explicitly here improves readability and
  4934 maintainability of proof texts.
  4935 
  4936 * Isar: improper proof element 'guess' is like 'obtain', but derives
  4937 the obtained context from the course of reasoning!  For example:
  4938 
  4939   assume "EX x y. A x & B y"   -- "any previous fact"
  4940   then guess x and y by clarify
  4941 
  4942 This technique is potentially adventurous, depending on the facts and
  4943 proof tools being involved here.
  4944 
  4945 * Isar: known facts from the proof context may be specified as literal
  4946 propositions, using ASCII back-quote syntax.  This works wherever
  4947 named facts used to be allowed so far, in proof commands, proof
  4948 methods, attributes etc.  Literal facts are retrieved from the context
  4949 according to unification of type and term parameters.  For example,
  4950 provided that "A" and "A ==> B" and "!!x. P x ==> Q x" are known
  4951 theorems in the current context, then these are valid literal facts:
  4952 `A` and `A ==> B` and `!!x. P x ==> Q x" as well as `P a ==> Q a` etc.
  4953 
  4954 There is also a proof method "fact" which does the same composition
  4955 for explicit goal states, e.g. the following proof texts coincide with
  4956 certain special cases of literal facts:
  4957 
  4958   have "A" by fact                 ==  note `A`
  4959   have "A ==> B" by fact           ==  note `A ==> B`
  4960   have "!!x. P x ==> Q x" by fact  ==  note `!!x. P x ==> Q x`
  4961   have "P a ==> Q a" by fact       ==  note `P a ==> Q a`
  4962 
  4963 * Isar: ":" (colon) is no longer a symbolic identifier character in
  4964 outer syntax.  Thus symbolic identifiers may be used without
  4965 additional white space in declarations like this: ``assume *: A''.
  4966 
  4967 * Isar: 'print_facts' prints all local facts of the current context,
  4968 both named and unnamed ones.
  4969 
  4970 * Isar: 'def' now admits simultaneous definitions, e.g.:
  4971 
  4972   def x == "t" and y == "u"
  4973 
  4974 * Isar: added command 'unfolding', which is structurally similar to
  4975 'using', but affects both the goal state and facts by unfolding given
  4976 rewrite rules.  Thus many occurrences of the 'unfold' method or
  4977 'unfolded' attribute may be replaced by first-class proof text.
  4978 
  4979 * Isar: methods 'unfold' / 'fold', attributes 'unfolded' / 'folded',
  4980 and command 'unfolding' now all support object-level equalities
  4981 (potentially conditional).  The underlying notion of rewrite rule is
  4982 analogous to the 'rule_format' attribute, but *not* that of the
  4983 Simplifier (which is usually more generous).
  4984 
  4985 * Isar: the new attribute [rotated n] (default n = 1) rotates the
  4986 premises of a theorem by n. Useful in conjunction with drule.
  4987 
  4988 * Isar: the goal restriction operator [N] (default N = 1) evaluates a
  4989 method expression within a sandbox consisting of the first N
  4990 sub-goals, which need to exist.  For example, ``simp_all [3]''
  4991 simplifies the first three sub-goals, while (rule foo, simp_all)[]
  4992 simplifies all new goals that emerge from applying rule foo to the
  4993 originally first one.
  4994 
  4995 * Isar: schematic goals are no longer restricted to higher-order
  4996 patterns; e.g. ``lemma "?P(?x)" by (rule TrueI)'' now works as
  4997 expected.
  4998 
  4999 * Isar: the conclusion of a long theorem statement is now either
  5000 'shows' (a simultaneous conjunction, as before), or 'obtains'
  5001 (essentially a disjunction of cases with local parameters and
  5002 assumptions).  The latter allows to express general elimination rules
  5003 adequately; in this notation common elimination rules look like this:
  5004 
  5005   lemma exE:    -- "EX x. P x ==> (!!x. P x ==> thesis) ==> thesis"
  5006     assumes "EX x. P x"
  5007     obtains x where "P x"
  5008 
  5009   lemma conjE:  -- "A & B ==> (A ==> B ==> thesis) ==> thesis"
  5010     assumes "A & B"
  5011     obtains A and B
  5012 
  5013   lemma disjE:  -- "A | B ==> (A ==> thesis) ==> (B ==> thesis) ==> thesis"
  5014     assumes "A | B"
  5015     obtains
  5016       A
  5017     | B
  5018 
  5019 The subsequent classical rules even refer to the formal "thesis"
  5020 explicitly:
  5021 
  5022   lemma classical:     -- "(~ thesis ==> thesis) ==> thesis"
  5023     obtains "~ thesis"
  5024 
  5025   lemma Peirce's_Law:  -- "((thesis ==> something) ==> thesis) ==> thesis"
  5026     obtains "thesis ==> something"
  5027 
  5028 The actual proof of an 'obtains' statement is analogous to that of the
  5029 Isar proof element 'obtain', only that there may be several cases.
  5030 Optional case names may be specified in parentheses; these will be
  5031 available both in the present proof and as annotations in the
  5032 resulting rule, for later use with the 'cases' method (cf. attribute
  5033 case_names).
  5034 
  5035 * Isar: the assumptions of a long theorem statement are available as
  5036 "assms" fact in the proof context.  This is more appropriate than the
  5037 (historical) "prems", which refers to all assumptions of the current
  5038 context, including those from the target locale, proof body etc.
  5039 
  5040 * Isar: 'print_statement' prints theorems from the current theory or
  5041 proof context in long statement form, according to the syntax of a
  5042 top-level lemma.
  5043 
  5044 * Isar: 'obtain' takes an optional case name for the local context
  5045 introduction rule (default "that").
  5046 
  5047 * Isar: removed obsolete 'concl is' patterns.  INCOMPATIBILITY, use
  5048 explicit (is "_ ==> ?foo") in the rare cases where this still happens
  5049 to occur.
  5050 
  5051 * Pure: syntax "CONST name" produces a fully internalized constant
  5052 according to the current context.  This is particularly useful for
  5053 syntax translations that should refer to internal constant
  5054 representations independently of name spaces.
  5055 
  5056 * Pure: syntax constant for foo (binder "FOO ") is called "foo_binder"
  5057 instead of "FOO ". This allows multiple binder declarations to coexist
  5058 in the same context.  INCOMPATIBILITY.
  5059 
  5060 * Isar/locales: 'notation' provides a robust interface to the 'syntax'
  5061 primitive that also works in a locale context (both for constants and
  5062 fixed variables). Type declaration and internal syntactic representation
  5063 of given constants retrieved from the context. Likewise, the
  5064 'no_notation' command allows to remove given syntax annotations from the
  5065 current context.
  5066 
  5067 * Isar/locales: new derived specification elements 'axiomatization',
  5068 'definition', 'abbreviation', which support type-inference, admit
  5069 object-level specifications (equality, equivalence).  See also the
  5070 isar-ref manual.  Examples:
  5071 
  5072   axiomatization
  5073     eq  (infix "===" 50) where
  5074     eq_refl: "x === x" and eq_subst: "x === y ==> P x ==> P y"
  5075 
  5076   definition "f x y = x + y + 1"
  5077   definition g where "g x = f x x"
  5078 
  5079   abbreviation
  5080     neq  (infix "=!=" 50) where
  5081     "x =!= y == ~ (x === y)"
  5082 
  5083 These specifications may be also used in a locale context.  Then the
  5084 constants being introduced depend on certain fixed parameters, and the
  5085 constant name is qualified by the locale base name.  An internal
  5086 abbreviation takes care for convenient input and output, making the
  5087 parameters implicit and using the original short name.  See also
  5088 src/HOL/ex/Abstract_NAT.thy for an example of deriving polymorphic
  5089 entities from a monomorphic theory.
  5090 
  5091 Presently, abbreviations are only available 'in' a target locale, but
  5092 not inherited by general import expressions.  Also note that
  5093 'abbreviation' may be used as a type-safe replacement for 'syntax' +
  5094 'translations' in common applications.  The "no_abbrevs" print mode
  5095 prevents folding of abbreviations in term output.
  5096 
  5097 Concrete syntax is attached to specified constants in internal form,
  5098 independently of name spaces.  The parse tree representation is
  5099 slightly different -- use 'notation' instead of raw 'syntax', and
  5100 'translations' with explicit "CONST" markup to accommodate this.
  5101 
  5102 * Pure/Isar: unified syntax for new-style specification mechanisms
  5103 (e.g.  'definition', 'abbreviation', or 'inductive' in HOL) admits
  5104 full type inference and dummy patterns ("_").  For example:
  5105 
  5106   definition "K x _ = x"
  5107 
  5108   inductive conj for A B
  5109   where "A ==> B ==> conj A B"
  5110 
  5111 * Pure: command 'print_abbrevs' prints all constant abbreviations of
  5112 the current context.  Print mode "no_abbrevs" prevents inversion of
  5113 abbreviations on output.
  5114 
  5115 * Isar/locales: improved parameter handling: use of locales "var" and
  5116 "struct" no longer necessary; - parameter renamings are no longer
  5117 required to be injective.  For example, this allows to define
  5118 endomorphisms as locale endom = homom mult mult h.
  5119 
  5120 * Isar/locales: changed the way locales with predicates are defined.
  5121 Instead of accumulating the specification, the imported expression is
  5122 now an interpretation.  INCOMPATIBILITY: different normal form of
  5123 locale expressions.  In particular, in interpretations of locales with
  5124 predicates, goals repesenting already interpreted fragments are not
  5125 removed automatically.  Use methods `intro_locales' and
  5126 `unfold_locales'; see below.
  5127 
  5128 * Isar/locales: new methods `intro_locales' and `unfold_locales'
  5129 provide backward reasoning on locales predicates.  The methods are
  5130 aware of interpretations and discharge corresponding goals.
  5131 `intro_locales' is less aggressive then `unfold_locales' and does not
  5132 unfold predicates to assumptions.
  5133 
  5134 * Isar/locales: the order in which locale fragments are accumulated
  5135 has changed.  This enables to override declarations from fragments due
  5136 to interpretations -- for example, unwanted simp rules.
  5137 
  5138 * Isar/locales: interpretation in theories and proof contexts has been
  5139 extended.  One may now specify (and prove) equations, which are
  5140 unfolded in interpreted theorems.  This is useful for replacing
  5141 defined concepts (constants depending on locale parameters) by
  5142 concepts already existing in the target context.  Example:
  5143 
  5144   interpretation partial_order ["op <= :: [int, int] => bool"]
  5145     where "partial_order.less (op <=) (x::int) y = (x < y)"
  5146 
  5147 Typically, the constant `partial_order.less' is created by a
  5148 definition specification element in the context of locale
  5149 partial_order.
  5150 
  5151 * Method "induct": improved internal context management to support
  5152 local fixes and defines on-the-fly. Thus explicit meta-level
  5153 connectives !!  and ==> are rarely required anymore in inductive goals
  5154 (using object-logic connectives for this purpose has been long
  5155 obsolete anyway). Common proof patterns are explained in
  5156 src/HOL/Induct/Common_Patterns.thy, see also
  5157 src/HOL/Isar_examples/Puzzle.thy and src/HOL/Lambda for realistic
  5158 examples.
  5159 
  5160 * Method "induct": improved handling of simultaneous goals. Instead of
  5161 introducing object-level conjunction, the statement is now split into
  5162 several conclusions, while the corresponding symbolic cases are nested
  5163 accordingly. INCOMPATIBILITY, proofs need to be structured explicitly,
  5164 see src/HOL/Induct/Common_Patterns.thy, for example.
  5165 
  5166 * Method "induct": mutual induction rules are now specified as a list
  5167 of rule sharing the same induction cases. HOL packages usually provide
  5168 foo_bar.inducts for mutually defined items foo and bar (e.g. inductive
  5169 predicates/sets or datatypes). INCOMPATIBILITY, users need to specify
  5170 mutual induction rules differently, i.e. like this:
  5171 
  5172   (induct rule: foo_bar.inducts)
  5173   (induct set: foo bar)
  5174   (induct pred: foo bar)
  5175   (induct type: foo bar)
  5176 
  5177 The ML function ProjectRule.projections turns old-style rules into the
  5178 new format.
  5179 
  5180 * Method "coinduct": dual of induction, see
  5181 src/HOL/Library/Coinductive_List.thy for various examples.
  5182 
  5183 * Method "cases", "induct", "coinduct": the ``(open)'' option is
  5184 considered a legacy feature.
  5185 
  5186 * Attribute "symmetric" produces result with standardized schematic
  5187 variables (index 0).  Potential INCOMPATIBILITY.
  5188 
  5189 * Simplifier: by default the simplifier trace only shows top level
  5190 rewrites now. That is, trace_simp_depth_limit is set to 1 by
  5191 default. Thus there is less danger of being flooded by the trace. The
  5192 trace indicates where parts have been suppressed.
  5193   
  5194 * Provers/classical: removed obsolete classical version of elim_format
  5195 attribute; classical elim/dest rules are now treated uniformly when
  5196 manipulating the claset.
  5197 
  5198 * Provers/classical: stricter checks to ensure that supplied intro,
  5199 dest and elim rules are well-formed; dest and elim rules must have at
  5200 least one premise.
  5201 
  5202 * Provers/classical: attributes dest/elim/intro take an optional
  5203 weight argument for the rule (just as the Pure versions).  Weights are
  5204 ignored by automated tools, but determine the search order of single
  5205 rule steps.
  5206 
  5207 * Syntax: input syntax now supports dummy variable binding "%_. b",
  5208 where the body does not mention the bound variable.  Note that dummy
  5209 patterns implicitly depend on their context of bounds, which makes
  5210 "{_. _}" match any set comprehension as expected.  Potential
  5211 INCOMPATIBILITY -- parse translations need to cope with syntactic
  5212 constant "_idtdummy" in the binding position.
  5213 
  5214 * Syntax: removed obsolete syntactic constant "_K" and its associated
  5215 parse translation.  INCOMPATIBILITY -- use dummy abstraction instead,
  5216 for example "A -> B" => "Pi A (%_. B)".
  5217 
  5218 * Pure: 'class_deps' command visualizes the subclass relation, using
  5219 the graph browser tool.
  5220 
  5221 * Pure: 'print_theory' now suppresses certain internal declarations by
  5222 default; use '!' option for full details.
  5223 
  5224 
  5225 *** HOL ***
  5226 
  5227 * Method "metis" proves goals by applying the Metis general-purpose
  5228 resolution prover (see also http://gilith.com/software/metis/).
  5229 Examples are in the directory MetisExamples.  WARNING: the
  5230 Isabelle/HOL-Metis integration does not yet work properly with
  5231 multi-threading.
  5232   
  5233 * Command 'sledgehammer' invokes external automatic theorem provers as
  5234 background processes.  It generates calls to the "metis" method if
  5235 successful. These can be pasted into the proof.  Users do not have to
  5236 wait for the automatic provers to return.  WARNING: does not really
  5237 work with multi-threading.
  5238 
  5239 * New "auto_quickcheck" feature tests outermost goal statements for
  5240 potential counter-examples.  Controlled by ML references
  5241 auto_quickcheck (default true) and auto_quickcheck_time_limit (default
  5242 5000 milliseconds).  Fails silently if statements is outside of
  5243 executable fragment, or any other codgenerator problem occurs.
  5244 
  5245 * New constant "undefined" with axiom "undefined x = undefined".
  5246 
  5247 * Added class "HOL.eq", allowing for code generation with polymorphic
  5248 equality.
  5249 
  5250 * Some renaming of class constants due to canonical name prefixing in
  5251 the new 'class' package:
  5252 
  5253     HOL.abs ~> HOL.abs_class.abs
  5254     HOL.divide ~> HOL.divide_class.divide
  5255     0 ~> HOL.zero_class.zero
  5256     1 ~> HOL.one_class.one
  5257     op + ~> HOL.plus_class.plus
  5258     op - ~> HOL.minus_class.minus
  5259     uminus ~> HOL.minus_class.uminus
  5260     op * ~> HOL.times_class.times
  5261     op < ~> HOL.ord_class.less
  5262     op <= > HOL.ord_class.less_eq
  5263     Nat.power ~> Power.power_class.power
  5264     Nat.size ~> Nat.size_class.size
  5265     Numeral.number_of ~> Numeral.number_class.number_of
  5266     FixedPoint.Inf ~> Lattices.complete_lattice_class.Inf
  5267     FixedPoint.Sup ~> Lattices.complete_lattice_class.Sup
  5268     Orderings.min ~> Orderings.ord_class.min
  5269     Orderings.max ~> Orderings.ord_class.max
  5270     Divides.op div ~> Divides.div_class.div
  5271     Divides.op mod ~> Divides.div_class.mod
  5272     Divides.op dvd ~> Divides.div_class.dvd
  5273 
  5274 INCOMPATIBILITY.  Adaptions may be required in the following cases:
  5275 
  5276 a) User-defined constants using any of the names "plus", "minus",
  5277 "times", "less" or "less_eq". The standard syntax translations for
  5278 "+", "-" and "*" may go wrong.  INCOMPATIBILITY: use more specific
  5279 names.
  5280 
  5281 b) Variables named "plus", "minus", "times", "less", "less_eq"
  5282 INCOMPATIBILITY: use more specific names.
  5283 
  5284 c) Permutative equations (e.g. "a + b = b + a")
  5285 Since the change of names also changes the order of terms, permutative
  5286 rewrite rules may get applied in a different order. Experience shows
  5287 that this is rarely the case (only two adaptions in the whole Isabelle
  5288 distribution).  INCOMPATIBILITY: rewrite proofs
  5289 
  5290 d) ML code directly refering to constant names
  5291 This in general only affects hand-written proof tactics, simprocs and
  5292 so on.  INCOMPATIBILITY: grep your sourcecode and replace names.
  5293 Consider using @{const_name} antiquotation.
  5294 
  5295 * New class "default" with associated constant "default".
  5296 
  5297 * Function "sgn" is now overloaded and available on int, real, complex
  5298 (and other numeric types), using class "sgn".  Two possible defs of
  5299 sgn are given as equational assumptions in the classes sgn_if and
  5300 sgn_div_norm; ordered_idom now also inherits from sgn_if.
  5301 INCOMPATIBILITY.
  5302 
  5303 * Locale "partial_order" now unified with class "order" (cf. theory
  5304 Orderings), added parameter "less".  INCOMPATIBILITY.
  5305 
  5306 * Renamings in classes "order" and "linorder": facts "refl", "trans" and
  5307 "cases" to "order_refl", "order_trans" and "linorder_cases", to avoid
  5308 clashes with HOL "refl" and "trans".  INCOMPATIBILITY.
  5309 
  5310 * Classes "order" and "linorder": potential INCOMPATIBILITY due to
  5311 changed order of proof goals in instance proofs.
  5312 
  5313 * The transitivity reasoner for partial and linear orders is set up
  5314 for classes "order" and "linorder".  Instances of the reasoner are available
  5315 in all contexts importing or interpreting the corresponding locales.
  5316 Method "order" invokes the reasoner separately; the reasoner
  5317 is also integrated with the Simplifier as a solver.  Diagnostic
  5318 command 'print_orders' shows the available instances of the reasoner
  5319 in the current context.
  5320 
  5321 * Localized monotonicity predicate in theory "Orderings"; integrated
  5322 lemmas max_of_mono and min_of_mono with this predicate.
  5323 INCOMPATIBILITY.
  5324 
  5325 * Formulation of theorem "dense" changed slightly due to integration
  5326 with new class dense_linear_order.
  5327 
  5328 * Uniform lattice theory development in HOL.
  5329 
  5330     constants "meet" and "join" now named "inf" and "sup"
  5331     constant "Meet" now named "Inf"
  5332 
  5333     classes "meet_semilorder" and "join_semilorder" now named
  5334       "lower_semilattice" and "upper_semilattice"
  5335     class "lorder" now named "lattice"
  5336     class "comp_lat" now named "complete_lattice"
  5337 
  5338     Instantiation of lattice classes allows explicit definitions
  5339     for "inf" and "sup" operations (or "Inf" and "Sup" for complete lattices).
  5340 
  5341   INCOMPATIBILITY.  Theorem renames:
  5342 
  5343     meet_left_le            ~> inf_le1
  5344     meet_right_le           ~> inf_le2
  5345     join_left_le            ~> sup_ge1
  5346     join_right_le           ~> sup_ge2
  5347     meet_join_le            ~> inf_sup_ord
  5348     le_meetI                ~> le_infI
  5349     join_leI                ~> le_supI
  5350     le_meet                 ~> le_inf_iff
  5351     le_join                 ~> ge_sup_conv
  5352     meet_idempotent         ~> inf_idem
  5353     join_idempotent         ~> sup_idem
  5354     meet_comm               ~> inf_commute
  5355     join_comm               ~> sup_commute
  5356     meet_leI1               ~> le_infI1
  5357     meet_leI2               ~> le_infI2
  5358     le_joinI1               ~> le_supI1
  5359     le_joinI2               ~> le_supI2
  5360     meet_assoc              ~> inf_assoc
  5361     join_assoc              ~> sup_assoc
  5362     meet_left_comm          ~> inf_left_commute
  5363     meet_left_idempotent    ~> inf_left_idem
  5364     join_left_comm          ~> sup_left_commute
  5365     join_left_idempotent    ~> sup_left_idem
  5366     meet_aci                ~> inf_aci
  5367     join_aci                ~> sup_aci
  5368     le_def_meet             ~> le_iff_inf
  5369     le_def_join             ~> le_iff_sup
  5370     join_absorp2            ~> sup_absorb2
  5371     join_absorp1            ~> sup_absorb1
  5372     meet_absorp1            ~> inf_absorb1
  5373     meet_absorp2            ~> inf_absorb2
  5374     meet_join_absorp        ~> inf_sup_absorb
  5375     join_meet_absorp        ~> sup_inf_absorb
  5376     distrib_join_le         ~> distrib_sup_le
  5377     distrib_meet_le         ~> distrib_inf_le
  5378 
  5379     add_meet_distrib_left   ~> add_inf_distrib_left
  5380     add_join_distrib_left   ~> add_sup_distrib_left
  5381     is_join_neg_meet        ~> is_join_neg_inf
  5382     is_meet_neg_join        ~> is_meet_neg_sup
  5383     add_meet_distrib_right  ~> add_inf_distrib_right
  5384     add_join_distrib_right  ~> add_sup_distrib_right
  5385     add_meet_join_distribs  ~> add_sup_inf_distribs
  5386     join_eq_neg_meet        ~> sup_eq_neg_inf
  5387     meet_eq_neg_join        ~> inf_eq_neg_sup
  5388     add_eq_meet_join        ~> add_eq_inf_sup
  5389     meet_0_imp_0            ~> inf_0_imp_0
  5390     join_0_imp_0            ~> sup_0_imp_0
  5391     meet_0_eq_0             ~> inf_0_eq_0
  5392     join_0_eq_0             ~> sup_0_eq_0
  5393     neg_meet_eq_join        ~> neg_inf_eq_sup
  5394     neg_join_eq_meet        ~> neg_sup_eq_inf
  5395     join_eq_if              ~> sup_eq_if
  5396 
  5397     mono_meet               ~> mono_inf
  5398     mono_join               ~> mono_sup
  5399     meet_bool_eq            ~> inf_bool_eq
  5400     join_bool_eq            ~> sup_bool_eq
  5401     meet_fun_eq             ~> inf_fun_eq
  5402     join_fun_eq             ~> sup_fun_eq
  5403     meet_set_eq             ~> inf_set_eq
  5404     join_set_eq             ~> sup_set_eq
  5405     meet1_iff               ~> inf1_iff
  5406     meet2_iff               ~> inf2_iff
  5407     meet1I                  ~> inf1I
  5408     meet2I                  ~> inf2I
  5409     meet1D1                 ~> inf1D1
  5410     meet2D1                 ~> inf2D1
  5411     meet1D2                 ~> inf1D2
  5412     meet2D2                 ~> inf2D2
  5413     meet1E                  ~> inf1E
  5414     meet2E                  ~> inf2E
  5415     join1_iff               ~> sup1_iff
  5416     join2_iff               ~> sup2_iff
  5417     join1I1                 ~> sup1I1
  5418     join2I1                 ~> sup2I1
  5419     join1I1                 ~> sup1I1
  5420     join2I2                 ~> sup1I2
  5421     join1CI                 ~> sup1CI
  5422     join2CI                 ~> sup2CI
  5423     join1E                  ~> sup1E
  5424     join2E                  ~> sup2E
  5425 
  5426     is_meet_Meet            ~> is_meet_Inf
  5427     Meet_bool_def           ~> Inf_bool_def
  5428     Meet_fun_def            ~> Inf_fun_def
  5429     Meet_greatest           ~> Inf_greatest
  5430     Meet_lower              ~> Inf_lower
  5431     Meet_set_def            ~> Inf_set_def
  5432 
  5433     Sup_def                 ~> Sup_Inf
  5434     Sup_bool_eq             ~> Sup_bool_def
  5435     Sup_fun_eq              ~> Sup_fun_def
  5436     Sup_set_eq              ~> Sup_set_def
  5437 
  5438     listsp_meetI            ~> listsp_infI
  5439     listsp_meet_eq          ~> listsp_inf_eq
  5440 
  5441     meet_min                ~> inf_min
  5442     join_max                ~> sup_max
  5443 
  5444 * Added syntactic class "size"; overloaded constant "size" now has
  5445 type "'a::size ==> bool"
  5446 
  5447 * Internal reorganisation of `size' of datatypes: size theorems
  5448 "foo.size" are no longer subsumed by "foo.simps" (but are still
  5449 simplification rules by default!); theorems "prod.size" now named
  5450 "*.size".
  5451 
  5452 * Class "div" now inherits from class "times" rather than "type".
  5453 INCOMPATIBILITY.
  5454 
  5455 * HOL/Finite_Set: "name-space" locales Lattice, Distrib_lattice,
  5456 Linorder etc.  have disappeared; operations defined in terms of
  5457 fold_set now are named Inf_fin, Sup_fin.  INCOMPATIBILITY.
  5458 
  5459 * HOL/Nat: neq0_conv no longer declared as iff.  INCOMPATIBILITY.
  5460 
  5461 * HOL-Word: New extensive library and type for generic, fixed size
  5462 machine words, with arithmetic, bit-wise, shifting and rotating
  5463 operations, reflection into int, nat, and bool lists, automation for
  5464 linear arithmetic (by automatic reflection into nat or int), including
  5465 lemmas on overflow and monotonicity.  Instantiated to all appropriate
  5466 arithmetic type classes, supporting automatic simplification of
  5467 numerals on all operations.
  5468 
  5469 * Library/Boolean_Algebra: locales for abstract boolean algebras.
  5470 
  5471 * Library/Numeral_Type: numbers as types, e.g. TYPE(32).
  5472 
  5473 * Code generator library theories:
  5474   - Code_Integer represents HOL integers by big integer literals in target
  5475     languages.
  5476   - Code_Char represents HOL characters by character literals in target
  5477     languages.
  5478   - Code_Char_chr like Code_Char, but also offers treatment of character
  5479     codes; includes Code_Integer.
  5480   - Executable_Set allows to generate code for finite sets using lists.
  5481   - Executable_Rat implements rational numbers as triples (sign, enumerator,
  5482     denominator).
  5483   - Executable_Real implements a subset of real numbers, namly those
  5484     representable by rational numbers.
  5485   - Efficient_Nat implements natural numbers by integers, which in general will
  5486     result in higher efficency; pattern matching with 0/Suc is eliminated;
  5487     includes Code_Integer.
  5488   - Code_Index provides an additional datatype index which is mapped to
  5489     target-language built-in integers.
  5490   - Code_Message provides an additional datatype message_string which is isomorphic to
  5491     strings; messages are mapped to target-language strings.
  5492 
  5493 * New package for inductive predicates
  5494 
  5495   An n-ary predicate p with m parameters z_1, ..., z_m can now be defined via
  5496 
  5497     inductive
  5498       p :: "U_1 => ... => U_m => T_1 => ... => T_n => bool"
  5499       for z_1 :: U_1 and ... and z_n :: U_m
  5500     where
  5501       rule_1: "... ==> p z_1 ... z_m t_1_1 ... t_1_n"
  5502     | ...
  5503 
  5504   with full support for type-inference, rather than
  5505 
  5506     consts s :: "U_1 => ... => U_m => (T_1 * ... * T_n) set"
  5507 
  5508     abbreviation p :: "U_1 => ... => U_m => T_1 => ... => T_n => bool"
  5509     where "p z_1 ... z_m x_1 ... x_n == (x_1, ..., x_n) : s z_1 ... z_m"
  5510 
  5511     inductive "s z_1 ... z_m"
  5512     intros
  5513       rule_1: "... ==> (t_1_1, ..., t_1_n) : s z_1 ... z_m"
  5514       ...
  5515 
  5516   For backward compatibility, there is a wrapper allowing inductive
  5517   sets to be defined with the new package via
  5518 
  5519     inductive_set
  5520       s :: "U_1 => ... => U_m => (T_1 * ... * T_n) set"
  5521       for z_1 :: U_1 and ... and z_n :: U_m
  5522     where
  5523       rule_1: "... ==> (t_1_1, ..., t_1_n) : s z_1 ... z_m"
  5524     | ...
  5525 
  5526   or
  5527 
  5528     inductive_set
  5529       s :: "U_1 => ... => U_m => (T_1 * ... * T_n) set"
  5530       and p :: "U_1 => ... => U_m => T_1 => ... => T_n => bool"
  5531       for z_1 :: U_1 and ... and z_n :: U_m
  5532     where
  5533       "p z_1 ... z_m x_1 ... x_n == (x_1, ..., x_n) : s z_1 ... z_m"
  5534     | rule_1: "... ==> p z_1 ... z_m t_1_1 ... t_1_n"
  5535     | ...
  5536 
  5537   if the additional syntax "p ..." is required.
  5538 
  5539   Numerous examples can be found in the subdirectories src/HOL/Auth,
  5540   src/HOL/Bali, src/HOL/Induct, and src/HOL/MicroJava.
  5541 
  5542   INCOMPATIBILITIES:
  5543 
  5544   - Since declaration and definition of inductive sets or predicates
  5545     is no longer separated, abbreviations involving the newly
  5546     introduced sets or predicates must be specified together with the
  5547     introduction rules after the 'where' keyword (see above), rather
  5548     than before the actual inductive definition.
  5549 
  5550   - The variables in induction and elimination rules are now
  5551     quantified in the order of their occurrence in the introduction
  5552     rules, rather than in alphabetical order. Since this may break
  5553     some proofs, these proofs either have to be repaired, e.g. by
  5554     reordering the variables a_i_1 ... a_i_{k_i} in Isar 'case'
  5555     statements of the form
  5556 
  5557       case (rule_i a_i_1 ... a_i_{k_i})
  5558 
  5559     or the old order of quantification has to be restored by explicitly adding
  5560     meta-level quantifiers in the introduction rules, i.e.
  5561 
  5562       | rule_i: "!!a_i_1 ... a_i_{k_i}. ... ==> p z_1 ... z_m t_i_1 ... t_i_n"
  5563 
  5564   - The format of the elimination rules is now
  5565 
  5566       p z_1 ... z_m x_1 ... x_n ==>
  5567         (!!a_1_1 ... a_1_{k_1}. x_1 = t_1_1 ==> ... ==> x_n = t_1_n ==> ... ==> P)
  5568         ==> ... ==> P
  5569 
  5570     for predicates and
  5571 
  5572       (x_1, ..., x_n) : s z_1 ... z_m ==>
  5573         (!!a_1_1 ... a_1_{k_1}. x_1 = t_1_1 ==> ... ==> x_n = t_1_n ==> ... ==> P)
  5574         ==> ... ==> P
  5575 
  5576     for sets rather than
  5577 
  5578       x : s z_1 ... z_m ==>
  5579         (!!a_1_1 ... a_1_{k_1}. x = (t_1_1, ..., t_1_n) ==> ... ==> P)
  5580         ==> ... ==> P
  5581 
  5582     This may require terms in goals to be expanded to n-tuples
  5583     (e.g. using case_tac or simplification with the split_paired_all
  5584     rule) before the above elimination rule is applicable.
  5585 
  5586   - The elimination or case analysis rules for (mutually) inductive
  5587     sets or predicates are now called "p_1.cases" ... "p_k.cases". The
  5588     list of rules "p_1_..._p_k.elims" is no longer available.
  5589 
  5590 * New package "function"/"fun" for general recursive functions,
  5591 supporting mutual and nested recursion, definitions in local contexts,
  5592 more general pattern matching and partiality. See HOL/ex/Fundefs.thy
  5593 for small examples, and the separate tutorial on the function
  5594 package. The old recdef "package" is still available as before, but
  5595 users are encouraged to use the new package.
  5596 
  5597 * Method "lexicographic_order" automatically synthesizes termination
  5598 relations as lexicographic combinations of size measures. 
  5599 
  5600 * Case-expressions allow arbitrary constructor-patterns (including
  5601 "_") and take their order into account, like in functional
  5602 programming.  Internally, this is translated into nested
  5603 case-expressions; missing cases are added and mapped to the predefined
  5604 constant "undefined". In complicated cases printing may no longer show
  5605 the original input but the internal form. Lambda-abstractions allow
  5606 the same form of pattern matching: "% pat1 => e1 | ..." is an
  5607 abbreviation for "%x. case x of pat1 => e1 | ..." where x is a new
  5608 variable.
  5609 
  5610 * IntDef: The constant "int :: nat => int" has been removed; now "int"
  5611 is an abbreviation for "of_nat :: nat => int". The simplification
  5612 rules for "of_nat" have been changed to work like "int" did
  5613 previously.  Potential INCOMPATIBILITY:
  5614   - "of_nat (Suc m)" simplifies to "1 + of_nat m" instead of "of_nat m + 1"
  5615   - of_nat_diff and of_nat_mult are no longer default simp rules
  5616 
  5617 * Method "algebra" solves polynomial equations over (semi)rings using
  5618 Groebner bases. The (semi)ring structure is defined by locales and the
  5619 tool setup depends on that generic context. Installing the method for
  5620 a specific type involves instantiating the locale and possibly adding
  5621 declarations for computation on the coefficients.  The method is
  5622 already instantiated for natural numbers and for the axiomatic class
  5623 of idoms with numerals.  See also the paper by Chaieb and Wenzel at
  5624 CALCULEMUS 2007 for the general principles underlying this
  5625 architecture of context-aware proof-tools.
  5626 
  5627 * Method "ferrack" implements quantifier elimination over
  5628 special-purpose dense linear orders using locales (analogous to
  5629 "algebra"). The method is already installed for class
  5630 {ordered_field,recpower,number_ring} which subsumes real, hyperreal,
  5631 rat, etc.
  5632 
  5633 * Former constant "List.op @" now named "List.append".  Use ML
  5634 antiquotations @{const_name List.append} or @{term " ... @ ... "} to
  5635 circumvent possible incompatibilities when working on ML level.
  5636 
  5637 * primrec: missing cases mapped to "undefined" instead of "arbitrary".
  5638 
  5639 * New function listsum :: 'a list => 'a for arbitrary monoids.
  5640 Special syntax: "SUM x <- xs. f x" (and latex variants)
  5641 
  5642 * New syntax for Haskell-like list comprehension (input only), eg.
  5643 [(x,y). x <- xs, y <- ys, x ~= y], see also src/HOL/List.thy.
  5644 
  5645 * The special syntax for function "filter" has changed from [x :
  5646 xs. P] to [x <- xs. P] to avoid an ambiguity caused by list
  5647 comprehension syntax, and for uniformity.  INCOMPATIBILITY.
  5648 
  5649 * [a..b] is now defined for arbitrary linear orders.  It used to be
  5650 defined on nat only, as an abbreviation for [a..<Suc b]
  5651 INCOMPATIBILITY.
  5652 
  5653 * Renamed lemma "set_take_whileD"  to "set_takeWhileD".
  5654 
  5655 * New functions "sorted" and "sort" in src/HOL/List.thy.
  5656 
  5657 * New lemma collection field_simps (an extension of ring_simps) for
  5658 manipulating (in)equations involving division. Multiplies with all
  5659 denominators that can be proved to be non-zero (in equations) or
  5660 positive/negative (in inequations).
  5661 
  5662 * Lemma collections ring_eq_simps, group_eq_simps and ring_distrib
  5663 have been improved and renamed to ring_simps, group_simps and
  5664 ring_distribs.  Removed lemmas field_xyz in theory Ring_and_Field
  5665 because they were subsumed by lemmas xyz.  INCOMPATIBILITY.
  5666 
  5667 * Theory Library/Commutative_Ring: switched from recdef to function
  5668 package; constants add, mul, pow now curried.  Infix syntax for
  5669 algebraic operations.
  5670 
  5671 * Dropped redundant lemma def_imp_eq in favor of meta_eq_to_obj_eq.
  5672 INCOMPATIBILITY.
  5673 
  5674 * Dropped redundant lemma if_def2 in favor of if_bool_eq_conj.
  5675 INCOMPATIBILITY.
  5676 
  5677 * HOL/records: generalised field-update to take a function on the
  5678 field rather than the new value: r(|A := x|) is translated to A_update
  5679 (K x) r The K-combinator that is internally used is called K_record.
  5680 INCOMPATIBILITY: Usage of the plain update functions has to be
  5681 adapted.
  5682  
  5683 * Class "semiring_0" now contains annihilation axioms x * 0 = 0 and 0
  5684 * x = 0, which are required for a semiring.  Richer structures do not
  5685 inherit from semiring_0 anymore, because this property is a theorem
  5686 there, not an axiom.  INCOMPATIBILITY: In instances of semiring_0,
  5687 there is more to prove, but this is mostly trivial.
  5688 
  5689 * Class "recpower" is generalized to arbitrary monoids, not just
  5690 commutative semirings.  INCOMPATIBILITY: may need to incorporate
  5691 commutativity or semiring properties additionally.
  5692 
  5693 * Constant "List.list_all2" in List.thy now uses authentic syntax.
  5694 INCOMPATIBILITY: translations containing list_all2 may go wrong,
  5695 better use 'abbreviation'.
  5696 
  5697 * Renamed constant "List.op mem" to "List.member".  INCOMPATIBILITY.
  5698 
  5699 * Numeral syntax: type 'bin' which was a mere type copy of 'int' has
  5700 been abandoned in favour of plain 'int'.  INCOMPATIBILITY --
  5701 significant changes for setting up numeral syntax for types:
  5702   - New constants Numeral.pred and Numeral.succ instead
  5703       of former Numeral.bin_pred and Numeral.bin_succ.
  5704   - Use integer operations instead of bin_add, bin_mult and so on.
  5705   - Numeral simplification theorems named Numeral.numeral_simps instead of Bin_simps.
  5706   - ML structure Bin_Simprocs now named Int_Numeral_Base_Simprocs.
  5707 
  5708 See src/HOL/Integ/IntArith.thy for an example setup.
  5709 
  5710 * Command 'normal_form' computes the normal form of a term that may
  5711 contain free variables.  For example ``normal_form "rev [a, b, c]"''
  5712 produces ``[b, c, a]'' (without proof).  This command is suitable for
  5713 heavy-duty computations because the functions are compiled to ML
  5714 first.  Correspondingly, a method "normalization" is provided.  See
  5715 further src/HOL/ex/NormalForm.thy and src/Tools/nbe.ML.
  5716 
  5717 * Alternative iff syntax "A <-> B" for equality on bool (with priority
  5718 25 like -->); output depends on the "iff" print_mode, the default is
  5719 "A = B" (with priority 50).
  5720 
  5721 * Relations less (<) and less_eq (<=) are also available on type bool.
  5722 Modified syntax to disallow nesting without explicit parentheses,
  5723 e.g. "(x < y) < z" or "x < (y < z)", but NOT "x < y < z".  Potential
  5724 INCOMPATIBILITY.
  5725 
  5726 * "LEAST x:A. P" expands to "LEAST x. x:A & P" (input only).
  5727 
  5728 * Relation composition operator "op O" now has precedence 75 and binds
  5729 stronger than union and intersection. INCOMPATIBILITY.
  5730 
  5731 * The old set interval syntax "{m..n(}" (and relatives) has been
  5732 removed.  Use "{m..<n}" (and relatives) instead.
  5733 
  5734 * In the context of the assumption "~(s = t)" the Simplifier rewrites
  5735 "t = s" to False (by simproc "neq").  INCOMPATIBILITY, consider using
  5736 ``declare [[simproc del: neq]]''.
  5737 
  5738 * Simplifier: "m dvd n" where m and n are numbers is evaluated to
  5739 True/False.
  5740 
  5741 * Theorem Cons_eq_map_conv no longer declared as "simp".
  5742 
  5743 * Theorem setsum_mult renamed to setsum_right_distrib.
  5744 
  5745 * Prefer ex1I over ex_ex1I in single-step reasoning, e.g. by the
  5746 ``rule'' method.
  5747 
  5748 * Reimplemented methods "sat" and "satx", with several improvements:
  5749 goals no longer need to be stated as "<prems> ==> False", equivalences
  5750 (i.e. "=" on type bool) are handled, variable names of the form
  5751 "lit_<n>" are no longer reserved, significant speedup.
  5752 
  5753 * Methods "sat" and "satx" can now replay MiniSat proof traces.
  5754 zChaff is still supported as well.
  5755 
  5756 * 'inductive' and 'datatype': provide projections of mutual rules,
  5757 bundled as foo_bar.inducts;
  5758 
  5759 * Library: moved theories Parity, GCD, Binomial, Infinite_Set to
  5760 Library.
  5761 
  5762 * Library: moved theory Accessible_Part to main HOL.
  5763 
  5764 * Library: added theory Coinductive_List of potentially infinite lists
  5765 as greatest fixed-point.
  5766 
  5767 * Library: added theory AssocList which implements (finite) maps as
  5768 association lists.
  5769 
  5770 * Method "evaluation" solves goals (i.e. a boolean expression)
  5771 efficiently by compiling it to ML.  The goal is "proved" (via an
  5772 oracle) if it evaluates to True.
  5773 
  5774 * Linear arithmetic now splits certain operators (e.g. min, max, abs)
  5775 also when invoked by the simplifier.  This results in the Simplifier
  5776 being more powerful on arithmetic goals.  INCOMPATIBILITY.
  5777 Configuration option fast_arith_split_limit=0 recovers the old
  5778 behavior.
  5779 
  5780 * Support for hex (0x20) and binary (0b1001) numerals.
  5781 
  5782 * New method: reify eqs (t), where eqs are equations for an
  5783 interpretation I :: 'a list => 'b => 'c and t::'c is an optional
  5784 parameter, computes a term s::'b and a list xs::'a list and proves the
  5785 theorem I xs s = t. This is also known as reification or quoting. The
  5786 resulting theorem is applied to the subgoal to substitute t with I xs
  5787 s.  If t is omitted, the subgoal itself is reified.
  5788 
  5789 * New method: reflection corr_thm eqs (t). The parameters eqs and (t)
  5790 are as explained above. corr_thm is a theorem for I vs (f t) = I vs t,
  5791 where f is supposed to be a computable function (in the sense of code
  5792 generattion). The method uses reify to compute s and xs as above then
  5793 applies corr_thm and uses normalization by evaluation to "prove" f s =
  5794 r and finally gets the theorem t = r, which is again applied to the
  5795 subgoal. An Example is available in src/HOL/ex/ReflectionEx.thy.
  5796 
  5797 * Reflection: Automatic reification now handels binding, an example is
  5798 available in src/HOL/ex/ReflectionEx.thy
  5799 
  5800 * HOL-Statespace: ``State Spaces: The Locale Way'' introduces a
  5801 command 'statespace' that is similar to 'record', but introduces an
  5802 abstract specification based on the locale infrastructure instead of
  5803 HOL types.  This leads to extra flexibility in composing state spaces,
  5804 in particular multiple inheritance and renaming of components.
  5805 
  5806 
  5807 *** HOL-Complex ***
  5808 
  5809 * Hyperreal: Functions root and sqrt are now defined on negative real
  5810 inputs so that root n (- x) = - root n x and sqrt (- x) = - sqrt x.
  5811 Nonnegativity side conditions have been removed from many lemmas, so
  5812 that more subgoals may now be solved by simplification; potential
  5813 INCOMPATIBILITY.
  5814 
  5815 * Real: new type classes formalize real normed vector spaces and
  5816 algebras, using new overloaded constants scaleR :: real => 'a => 'a
  5817 and norm :: 'a => real.
  5818 
  5819 * Real: constant of_real :: real => 'a::real_algebra_1 injects from
  5820 reals into other types. The overloaded constant Reals :: 'a set is now
  5821 defined as range of_real; potential INCOMPATIBILITY.
  5822 
  5823 * Real: proper support for ML code generation, including 'quickcheck'.
  5824 Reals are implemented as arbitrary precision rationals.
  5825 
  5826 * Hyperreal: Several constants that previously worked only for the
  5827 reals have been generalized, so they now work over arbitrary vector
  5828 spaces. Type annotations may need to be added in some cases; potential
  5829 INCOMPATIBILITY.
  5830 
  5831   Infinitesimal  :: ('a::real_normed_vector) star set
  5832   HFinite        :: ('a::real_normed_vector) star set
  5833   HInfinite      :: ('a::real_normed_vector) star set
  5834   approx         :: ('a::real_normed_vector) star => 'a star => bool
  5835   monad          :: ('a::real_normed_vector) star => 'a star set
  5836   galaxy         :: ('a::real_normed_vector) star => 'a star set
  5837   (NS)LIMSEQ     :: [nat => 'a::real_normed_vector, 'a] => bool
  5838   (NS)convergent :: (nat => 'a::real_normed_vector) => bool
  5839   (NS)Bseq       :: (nat => 'a::real_normed_vector) => bool
  5840   (NS)Cauchy     :: (nat => 'a::real_normed_vector) => bool
  5841   (NS)LIM        :: ['a::real_normed_vector => 'b::real_normed_vector, 'a, 'b] => bool
  5842   is(NS)Cont     :: ['a::real_normed_vector => 'b::real_normed_vector, 'a] => bool
  5843   deriv          :: ['a::real_normed_field => 'a, 'a, 'a] => bool
  5844   sgn            :: 'a::real_normed_vector => 'a
  5845   exp            :: 'a::{recpower,real_normed_field,banach} => 'a
  5846 
  5847 * Complex: Some complex-specific constants are now abbreviations for
  5848 overloaded ones: complex_of_real = of_real, cmod = norm, hcmod =
  5849 hnorm.  Other constants have been entirely removed in favor of the
  5850 polymorphic versions (INCOMPATIBILITY):
  5851 
  5852   approx        <-- capprox
  5853   HFinite       <-- CFinite
  5854   HInfinite     <-- CInfinite
  5855   Infinitesimal <-- CInfinitesimal
  5856   monad         <-- cmonad
  5857   galaxy        <-- cgalaxy
  5858   (NS)LIM       <-- (NS)CLIM, (NS)CRLIM
  5859   is(NS)Cont    <-- is(NS)Contc, is(NS)contCR
  5860   (ns)deriv     <-- (ns)cderiv
  5861 
  5862 
  5863 *** HOL-Algebra ***
  5864 
  5865 * Formalisation of ideals and the quotient construction over rings.
  5866 
  5867 * Order and lattice theory no longer based on records.
  5868 INCOMPATIBILITY.
  5869 
  5870 * Renamed lemmas least_carrier -> least_closed and greatest_carrier ->
  5871 greatest_closed.  INCOMPATIBILITY.
  5872 
  5873 * Method algebra is now set up via an attribute.  For examples see
  5874 Ring.thy.  INCOMPATIBILITY: the method is now weaker on combinations
  5875 of algebraic structures.
  5876 
  5877 * Renamed theory CRing to Ring.
  5878 
  5879 
  5880 *** HOL-Nominal ***
  5881 
  5882 * Substantial, yet incomplete support for nominal datatypes (binding
  5883 structures) based on HOL-Nominal logic.  See src/HOL/Nominal and
  5884 src/HOL/Nominal/Examples.  Prospective users should consult
  5885 http://isabelle.in.tum.de/nominal/
  5886 
  5887 
  5888 *** ML ***
  5889 
  5890 * ML basics: just one true type int, which coincides with IntInf.int
  5891 (even on SML/NJ).
  5892 
  5893 * ML within Isar: antiquotations allow to embed statically-checked
  5894 formal entities in the source, referring to the context available at
  5895 compile-time.  For example:
  5896 
  5897 ML {* @{sort "{zero,one}"} *}
  5898 ML {* @{typ "'a => 'b"} *}
  5899 ML {* @{term "%x. x"} *}
  5900 ML {* @{prop "x == y"} *}
  5901 ML {* @{ctyp "'a => 'b"} *}
  5902 ML {* @{cterm "%x. x"} *}
  5903 ML {* @{cprop "x == y"} *}
  5904 ML {* @{thm asm_rl} *}
  5905 ML {* @{thms asm_rl} *}
  5906 ML {* @{type_name c} *}
  5907 ML {* @{type_syntax c} *}
  5908 ML {* @{const_name c} *}
  5909 ML {* @{const_syntax c} *}
  5910 ML {* @{context} *}
  5911 ML {* @{theory} *}
  5912 ML {* @{theory Pure} *}
  5913 ML {* @{theory_ref} *}
  5914 ML {* @{theory_ref Pure} *}
  5915 ML {* @{simpset} *}
  5916 ML {* @{claset} *}
  5917 ML {* @{clasimpset} *}
  5918 
  5919 The same works for sources being ``used'' within an Isar context.
  5920 
  5921 * ML in Isar: improved error reporting; extra verbosity with
  5922 ML_Context.trace enabled.
  5923 
  5924 * Pure/General/table.ML: the join operations now works via exceptions
  5925 DUP/SAME instead of type option. This is simpler in simple cases, and
  5926 admits slightly more efficient complex applications.
  5927 
  5928 * Pure: 'advanced' translation functions (parse_translation etc.) now
  5929 use Context.generic instead of just theory.
  5930 
  5931 * Pure: datatype Context.generic joins theory/Proof.context and
  5932 provides some facilities for code that works in either kind of
  5933 context, notably GenericDataFun for uniform theory and proof data.
  5934 
  5935 * Pure: simplified internal attribute type, which is now always
  5936 Context.generic * thm -> Context.generic * thm. Global (theory) vs.
  5937 local (Proof.context) attributes have been discontinued, while
  5938 minimizing code duplication. Thm.rule_attribute and
  5939 Thm.declaration_attribute build canonical attributes; see also structure
  5940 Context for further operations on Context.generic, notably
  5941 GenericDataFun. INCOMPATIBILITY, need to adapt attribute type
  5942 declarations and definitions.
  5943 
  5944 * Context data interfaces (Theory/Proof/GenericDataFun): removed
  5945 name/print, uninitialized data defaults to ad-hoc copy of empty value,
  5946 init only required for impure data. INCOMPATIBILITY: empty really need
  5947 to be empty (no dependencies on theory content!)
  5948 
  5949 * Pure/kernel: consts certification ignores sort constraints given in
  5950 signature declarations. (This information is not relevant to the
  5951 logic, but only for type inference.) SIGNIFICANT INTERNAL CHANGE,
  5952 potential INCOMPATIBILITY.
  5953 
  5954 * Pure: axiomatic type classes are now purely definitional, with
  5955 explicit proofs of class axioms and super class relations performed
  5956 internally. See Pure/axclass.ML for the main internal interfaces --
  5957 notably AxClass.define_class supercedes AxClass.add_axclass, and
  5958 AxClass.axiomatize_class/classrel/arity supersede
  5959 Sign.add_classes/classrel/arities.
  5960 
  5961 * Pure/Isar: Args/Attrib parsers operate on Context.generic --
  5962 global/local versions on theory vs. Proof.context have been
  5963 discontinued; Attrib.syntax and Method.syntax have been adapted
  5964 accordingly.  INCOMPATIBILITY, need to adapt parser expressions for
  5965 attributes, methods, etc.
  5966 
  5967 * Pure: several functions of signature "... -> theory -> theory * ..."
  5968 have been reoriented to "... -> theory -> ... * theory" in order to
  5969 allow natural usage in combination with the ||>, ||>>, |-> and
  5970 fold_map combinators.
  5971 
  5972 * Pure: official theorem names (closed derivations) and additional
  5973 comments (tags) are now strictly separate.  Name hints -- which are
  5974 maintained as tags -- may be attached any time without affecting the
  5975 derivation.
  5976 
  5977 * Pure: primitive rule lift_rule now takes goal cterm instead of an
  5978 actual goal state (thm).  Use Thm.lift_rule (Thm.cprem_of st i) to
  5979 achieve the old behaviour.
  5980 
  5981 * Pure: the "Goal" constant is now called "prop", supporting a
  5982 slightly more general idea of ``protecting'' meta-level rule
  5983 statements.
  5984 
  5985 * Pure: Logic.(un)varify only works in a global context, which is now
  5986 enforced instead of silently assumed.  INCOMPATIBILITY, may use
  5987 Logic.legacy_(un)varify as temporary workaround.
  5988 
  5989 * Pure: structure Name provides scalable operations for generating
  5990 internal variable names, notably Name.variants etc.  This replaces
  5991 some popular functions from term.ML:
  5992 
  5993   Term.variant		->  Name.variant
  5994   Term.variantlist	->  Name.variant_list
  5995   Term.invent_names	->  Name.invent_list
  5996 
  5997 Note that low-level renaming rarely occurs in new code -- operations
  5998 from structure Variable are used instead (see below).
  5999 
  6000 * Pure: structure Variable provides fundamental operations for proper
  6001 treatment of fixed/schematic variables in a context.  For example,
  6002 Variable.import introduces fixes for schematics of given facts and
  6003 Variable.export reverses the effect (up to renaming) -- this replaces
  6004 various freeze_thaw operations.
  6005 
  6006 * Pure: structure Goal provides simple interfaces for
  6007 init/conclude/finish and tactical prove operations (replacing former
  6008 Tactic.prove).  Goal.prove is the canonical way to prove results
  6009 within a given context; Goal.prove_global is a degraded version for
  6010 theory level goals, including a global Drule.standard.  Note that
  6011 OldGoals.prove_goalw_cterm has long been obsolete, since it is
  6012 ill-behaved in a local proof context (e.g. with local fixes/assumes or
  6013 in a locale context).
  6014 
  6015 * Pure/Syntax: generic interfaces for parsing (Syntax.parse_term etc.)
  6016 and type checking (Syntax.check_term etc.), with common combinations
  6017 (Syntax.read_term etc.). These supersede former Sign.read_term etc.
  6018 which are considered legacy and await removal.
  6019 
  6020 * Pure/Syntax: generic interfaces for type unchecking
  6021 (Syntax.uncheck_terms etc.) and unparsing (Syntax.unparse_term etc.),
  6022 with common combinations (Syntax.pretty_term, Syntax.string_of_term
  6023 etc.).  Former Sign.pretty_term, Sign.string_of_term etc. are still
  6024 available for convenience, but refer to the very same operations using
  6025 a mere theory instead of a full context.
  6026 
  6027 * Isar: simplified treatment of user-level errors, using exception
  6028 ERROR of string uniformly.  Function error now merely raises ERROR,
  6029 without any side effect on output channels.  The Isar toplevel takes
  6030 care of proper display of ERROR exceptions.  ML code may use plain
  6031 handle/can/try; cat_error may be used to concatenate errors like this:
  6032 
  6033   ... handle ERROR msg => cat_error msg "..."
  6034 
  6035 Toplevel ML code (run directly or through the Isar toplevel) may be
  6036 embedded into the Isar toplevel with exception display/debug like
  6037 this:
  6038 
  6039   Isar.toplevel (fn () => ...)
  6040 
  6041 INCOMPATIBILITY, removed special transform_error facilities, removed
  6042 obsolete variants of user-level exceptions (ERROR_MESSAGE,
  6043 Context.PROOF, ProofContext.CONTEXT, Proof.STATE, ProofHistory.FAIL)
  6044 -- use plain ERROR instead.
  6045 
  6046 * Isar: theory setup now has type (theory -> theory), instead of a
  6047 list.  INCOMPATIBILITY, may use #> to compose setup functions.
  6048 
  6049 * Isar: ML toplevel pretty printer for type Proof.context, subject to
  6050 ProofContext.debug/verbose flags.
  6051 
  6052 * Isar: Toplevel.theory_to_proof admits transactions that modify the
  6053 theory before entering a proof state.  Transactions now always see a
  6054 quasi-functional intermediate checkpoint, both in interactive and
  6055 batch mode.
  6056 
  6057 * Isar: simplified interfaces for outer syntax.  Renamed
  6058 OuterSyntax.add_keywords to OuterSyntax.keywords.  Removed
  6059 OuterSyntax.add_parsers -- this functionality is now included in
  6060 OuterSyntax.command etc.  INCOMPATIBILITY.
  6061 
  6062 * Simplifier: the simpset of a running simplification process now
  6063 contains a proof context (cf. Simplifier.the_context), which is the
  6064 very context that the initial simpset has been retrieved from (by
  6065 simpset_of/local_simpset_of).  Consequently, all plug-in components
  6066 (solver, looper etc.) may depend on arbitrary proof data.
  6067 
  6068 * Simplifier.inherit_context inherits the proof context (plus the
  6069 local bounds) of the current simplification process; any simproc
  6070 etc. that calls the Simplifier recursively should do this!  Removed
  6071 former Simplifier.inherit_bounds, which is already included here --
  6072 INCOMPATIBILITY.  Tools based on low-level rewriting may even have to
  6073 specify an explicit context using Simplifier.context/theory_context.
  6074 
  6075 * Simplifier/Classical Reasoner: more abstract interfaces
  6076 change_simpset/claset for modifying the simpset/claset reference of a
  6077 theory; raw versions simpset/claset_ref etc. have been discontinued --
  6078 INCOMPATIBILITY.
  6079 
  6080 * Provers: more generic wrt. syntax of object-logics, avoid hardwired
  6081 "Trueprop" etc.
  6082 
  6083 
  6084 *** System ***
  6085 
  6086 * settings: the default heap location within ISABELLE_HOME_USER now
  6087 includes ISABELLE_IDENTIFIER.  This simplifies use of multiple
  6088 Isabelle installations.
  6089 
  6090 * isabelle-process: option -S (secure mode) disables some critical
  6091 operations, notably runtime compilation and evaluation of ML source
  6092 code.
  6093 
  6094 * Basic Isabelle mode for jEdit, see Isabelle/lib/jedit/.
  6095 
  6096 * Support for parallel execution, using native multicore support of
  6097 Poly/ML 5.1.  The theory loader exploits parallelism when processing
  6098 independent theories, according to the given theory header
  6099 specifications. The maximum number of worker threads is specified via
  6100 usedir option -M or the "max-threads" setting in Proof General. A
  6101 speedup factor of 1.5--3.5 can be expected on a 4-core machine, and up
  6102 to 6 on a 8-core machine.  User-code needs to observe certain
  6103 guidelines for thread-safe programming, see appendix A in the Isar
  6104 Implementation manual.
  6105 
  6106 
  6107 
  6108 New in Isabelle2005 (October 2005)
  6109 ----------------------------------
  6110 
  6111 *** General ***
  6112 
  6113 * Theory headers: the new header syntax for Isar theories is
  6114 
  6115   theory <name>
  6116   imports <theory1> ... <theoryN>
  6117   uses <file1> ... <fileM>
  6118   begin
  6119 
  6120 where the 'uses' part is optional.  The previous syntax
  6121 
  6122   theory <name> = <theory1> + ... + <theoryN>:
  6123 
  6124 will disappear in the next release.  Use isatool fixheaders to convert
  6125 existing theory files.  Note that there is no change in ancient
  6126 non-Isar theories now, but these will disappear soon.
  6127 
  6128 * Theory loader: parent theories can now also be referred to via
  6129 relative and absolute paths.
  6130 
  6131 * Command 'find_theorems' searches for a list of criteria instead of a
  6132 list of constants. Known criteria are: intro, elim, dest, name:string,
  6133 simp:term, and any term. Criteria can be preceded by '-' to select
  6134 theorems that do not match. Intro, elim, dest select theorems that
  6135 match the current goal, name:s selects theorems whose fully qualified
  6136 name contain s, and simp:term selects all simplification rules whose
  6137 lhs match term.  Any other term is interpreted as pattern and selects
  6138 all theorems matching the pattern. Available in ProofGeneral under
  6139 'ProofGeneral -> Find Theorems' or C-c C-f.  Example:
  6140 
  6141   C-c C-f (100) "(_::nat) + _ + _" intro -name: "HOL."
  6142 
  6143 prints the last 100 theorems matching the pattern "(_::nat) + _ + _",
  6144 matching the current goal as introduction rule and not having "HOL."
  6145 in their name (i.e. not being defined in theory HOL).
  6146 
  6147 * Command 'thms_containing' has been discontinued in favour of
  6148 'find_theorems'; INCOMPATIBILITY.
  6149 
  6150 * Communication with Proof General is now 8bit clean, which means that
  6151 Unicode text in UTF-8 encoding may be used within theory texts (both
  6152 formal and informal parts).  Cf. option -U of the Isabelle Proof
  6153 General interface.  Here are some simple examples (cf. src/HOL/ex):
  6154 
  6155   http://isabelle.in.tum.de/library/HOL/ex/Hebrew.html
  6156   http://isabelle.in.tum.de/library/HOL/ex/Chinese.html
  6157 
  6158 * Improved efficiency of the Simplifier and, to a lesser degree, the
  6159 Classical Reasoner.  Typical big applications run around 2 times
  6160 faster.
  6161 
  6162 
  6163 *** Document preparation ***
  6164 
  6165 * Commands 'display_drafts' and 'print_drafts' perform simple output
  6166 of raw sources.  Only those symbols that do not require additional
  6167 LaTeX packages (depending on comments in isabellesym.sty) are
  6168 displayed properly, everything else is left verbatim.  isatool display
  6169 and isatool print are used as front ends (these are subject to the
  6170 DVI/PDF_VIEWER and PRINT_COMMAND settings, respectively).
  6171 
  6172 * Command tags control specific markup of certain regions of text,
  6173 notably folding and hiding.  Predefined tags include "theory" (for
  6174 theory begin and end), "proof" for proof commands, and "ML" for
  6175 commands involving ML code; the additional tags "visible" and
  6176 "invisible" are unused by default.  Users may give explicit tag
  6177 specifications in the text, e.g. ''by %invisible (auto)''.  The
  6178 interpretation of tags is determined by the LaTeX job during document
  6179 preparation: see option -V of isatool usedir, or options -n and -t of
  6180 isatool document, or even the LaTeX macros \isakeeptag, \isafoldtag,
  6181 \isadroptag.
  6182 
  6183 Several document versions may be produced at the same time via isatool
  6184 usedir (the generated index.html will link all of them).  Typical
  6185 specifications include ''-V document=theory,proof,ML'' to present
  6186 theory/proof/ML parts faithfully, ''-V outline=/proof,/ML'' to fold
  6187 proof and ML commands, and ''-V mutilated=-theory,-proof,-ML'' to omit
  6188 these parts without any formal replacement text.  The Isabelle site
  6189 default settings produce ''document'' and ''outline'' versions as
  6190 specified above.
  6191 
  6192 * Several new antiquotations:
  6193 
  6194   @{term_type term} prints a term with its type annotated;
  6195 
  6196   @{typeof term} prints the type of a term;
  6197 
  6198   @{const const} is the same as @{term const}, but checks that the
  6199   argument is a known logical constant;
  6200 
  6201   @{term_style style term} and @{thm_style style thm} print a term or
  6202   theorem applying a "style" to it
  6203 
  6204   @{ML text}
  6205 
  6206 Predefined styles are 'lhs' and 'rhs' printing the lhs/rhs of
  6207 definitions, equations, inequations etc., 'concl' printing only the
  6208 conclusion of a meta-logical statement theorem, and 'prem1' .. 'prem19'
  6209 to print the specified premise.  TermStyle.add_style provides an ML
  6210 interface for introducing further styles.  See also the "LaTeX Sugar"
  6211 document practical applications.  The ML antiquotation prints
  6212 type-checked ML expressions verbatim.
  6213 
  6214 * Markup commands 'chapter', 'section', 'subsection', 'subsubsection',
  6215 and 'text' support optional locale specification '(in loc)', which
  6216 specifies the default context for interpreting antiquotations.  For
  6217 example: 'text (in lattice) {* @{thm inf_assoc}*}'.
  6218 
  6219 * Option 'locale=NAME' of antiquotations specifies an alternative
  6220 context interpreting the subsequent argument.  For example: @{thm
  6221 [locale=lattice] inf_assoc}.
  6222 
  6223 * Proper output of proof terms (@{prf ...} and @{full_prf ...}) within
  6224 a proof context.
  6225 
  6226 * Proper output of antiquotations for theory commands involving a
  6227 proof context (such as 'locale' or 'theorem (in loc) ...').
  6228 
  6229 * Delimiters of outer tokens (string etc.) now produce separate LaTeX
  6230 macros (\isachardoublequoteopen, isachardoublequoteclose etc.).
  6231 
  6232 * isatool usedir: new option -C (default true) controls whether option
  6233 -D should include a copy of the original document directory; -C false
  6234 prevents unwanted effects such as copying of administrative CVS data.
  6235 
  6236 
  6237 *** Pure ***
  6238 
  6239 * Considerably improved version of 'constdefs' command.  Now performs
  6240 automatic type-inference of declared constants; additional support for
  6241 local structure declarations (cf. locales and HOL records), see also
  6242 isar-ref manual.  Potential INCOMPATIBILITY: need to observe strictly
  6243 sequential dependencies of definitions within a single 'constdefs'
  6244 section; moreover, the declared name needs to be an identifier.  If
  6245 all fails, consider to fall back on 'consts' and 'defs' separately.
  6246 
  6247 * Improved indexed syntax and implicit structures.  First of all,
  6248 indexed syntax provides a notational device for subscripted
  6249 application, using the new syntax \<^bsub>term\<^esub> for arbitrary
  6250 expressions.  Secondly, in a local context with structure
  6251 declarations, number indexes \<^sub>n or the empty index (default
  6252 number 1) refer to a certain fixed variable implicitly; option
  6253 show_structs controls printing of implicit structures.  Typical
  6254 applications of these concepts involve record types and locales.
  6255 
  6256 * New command 'no_syntax' removes grammar declarations (and
  6257 translations) resulting from the given syntax specification, which is
  6258 interpreted in the same manner as for the 'syntax' command.
  6259 
  6260 * 'Advanced' translation functions (parse_translation etc.) may depend
  6261 on the signature of the theory context being presently used for
  6262 parsing/printing, see also isar-ref manual.
  6263 
  6264 * Improved 'oracle' command provides a type-safe interface to turn an
  6265 ML expression of type theory -> T -> term into a primitive rule of
  6266 type theory -> T -> thm (i.e. the functionality of Thm.invoke_oracle
  6267 is already included here); see also FOL/ex/IffExample.thy;
  6268 INCOMPATIBILITY.
  6269 
  6270 * axclass: name space prefix for class "c" is now "c_class" (was "c"
  6271 before); "cI" is no longer bound, use "c.intro" instead.
  6272 INCOMPATIBILITY.  This change avoids clashes of fact bindings for
  6273 axclasses vs. locales.
  6274 
  6275 * Improved internal renaming of symbolic identifiers -- attach primes
  6276 instead of base 26 numbers.
  6277 
  6278 * New flag show_question_marks controls printing of leading question
  6279 marks in schematic variable names.
  6280 
  6281 * In schematic variable names, *any* symbol following \<^isub> or
  6282 \<^isup> is now treated as part of the base name.  For example, the
  6283 following works without printing of awkward ".0" indexes:
  6284 
  6285   lemma "x\<^isub>1 = x\<^isub>2 ==> x\<^isub>2 = x\<^isub>1"
  6286     by simp
  6287 
  6288 * Inner syntax includes (*(*nested*) comments*).
  6289 
  6290 * Pretty printer now supports unbreakable blocks, specified in mixfix
  6291 annotations as "(00...)".
  6292 
  6293 * Clear separation of logical types and nonterminals, where the latter
  6294 may only occur in 'syntax' specifications or type abbreviations.
  6295 Before that distinction was only partially implemented via type class
  6296 "logic" vs. "{}".  Potential INCOMPATIBILITY in rare cases of improper
  6297 use of 'types'/'consts' instead of 'nonterminals'/'syntax'.  Some very
  6298 exotic syntax specifications may require further adaption
  6299 (e.g. Cube/Cube.thy).
  6300 
  6301 * Removed obsolete type class "logic", use the top sort {} instead.
  6302 Note that non-logical types should be declared as 'nonterminals'
  6303 rather than 'types'.  INCOMPATIBILITY for new object-logic
  6304 specifications.
  6305 
  6306 * Attributes 'induct' and 'cases': type or set names may now be
  6307 locally fixed variables as well.
  6308 
  6309 * Simplifier: can now control the depth to which conditional rewriting
  6310 is traced via the PG menu Isabelle -> Settings -> Trace Simp Depth
  6311 Limit.
  6312 
  6313 * Simplifier: simplification procedures may now take the current
  6314 simpset into account (cf. Simplifier.simproc(_i) / mk_simproc
  6315 interface), which is very useful for calling the Simplifier
  6316 recursively.  Minor INCOMPATIBILITY: the 'prems' argument of simprocs
  6317 is gone -- use prems_of_ss on the simpset instead.  Moreover, the
  6318 low-level mk_simproc no longer applies Logic.varify internally, to
  6319 allow for use in a context of fixed variables.
  6320 
  6321 * thin_tac now works even if the assumption being deleted contains !!
  6322 or ==>.  More generally, erule now works even if the major premise of
  6323 the elimination rule contains !! or ==>.
  6324 
  6325 * Method 'rules' has been renamed to 'iprover'. INCOMPATIBILITY.
  6326 
  6327 * Reorganized bootstrapping of the Pure theories; CPure is now derived
  6328 from Pure, which contains all common declarations already.  Both
  6329 theories are defined via plain Isabelle/Isar .thy files.
  6330 INCOMPATIBILITY: elements of CPure (such as the CPure.intro /
  6331 CPure.elim / CPure.dest attributes) now appear in the Pure name space;
  6332 use isatool fixcpure to adapt your theory and ML sources.
  6333 
  6334 * New syntax 'name(i-j, i-, i, ...)' for referring to specific
  6335 selections of theorems in named facts via index ranges.
  6336 
  6337 * 'print_theorems': in theory mode, really print the difference
  6338 wrt. the last state (works for interactive theory development only),
  6339 in proof mode print all local facts (cf. 'print_facts');
  6340 
  6341 * 'hide': option '(open)' hides only base names.
  6342 
  6343 * More efficient treatment of intermediate checkpoints in interactive
  6344 theory development.
  6345 
  6346 * Code generator is now invoked via code_module (incremental code
  6347 generation) and code_library (modular code generation, ML structures
  6348 for each theory).  INCOMPATIBILITY: new keywords 'file' and 'contains'
  6349 must be quoted when used as identifiers.
  6350 
  6351 * New 'value' command for reading, evaluating and printing terms using
  6352 the code generator.  INCOMPATIBILITY: command keyword 'value' must be
  6353 quoted when used as identifier.
  6354 
  6355 
  6356 *** Locales ***
  6357 
  6358 * New commands for the interpretation of locale expressions in
  6359 theories (1), locales (2) and proof contexts (3).  These generate
  6360 proof obligations from the expression specification.  After the
  6361 obligations have been discharged, theorems of the expression are added
  6362 to the theory, target locale or proof context.  The synopsis of the
  6363 commands is a follows:
  6364 
  6365   (1) interpretation expr inst
  6366   (2) interpretation target < expr
  6367   (3) interpret expr inst
  6368 
  6369 Interpretation in theories and proof contexts require a parameter
  6370 instantiation of terms from the current context.  This is applied to
  6371 specifications and theorems of the interpreted expression.
  6372 Interpretation in locales only permits parameter renaming through the
  6373 locale expression.  Interpretation is smart in that interpretations
  6374 that are active already do not occur in proof obligations, neither are
  6375 instantiated theorems stored in duplicate.  Use 'print_interps' to
  6376 inspect active interpretations of a particular locale.  For details,
  6377 see the Isar Reference manual.  Examples can be found in
  6378 HOL/Finite_Set.thy and HOL/Algebra/UnivPoly.thy.
  6379 
  6380 INCOMPATIBILITY: former 'instantiate' has been withdrawn, use
  6381 'interpret' instead.
  6382 
  6383 * New context element 'constrains' for adding type constraints to
  6384 parameters.
  6385 
  6386 * Context expressions: renaming of parameters with syntax
  6387 redeclaration.
  6388 
  6389 * Locale declaration: 'includes' disallowed.
  6390 
  6391 * Proper static binding of attribute syntax -- i.e. types / terms /
  6392 facts mentioned as arguments are always those of the locale definition
  6393 context, independently of the context of later invocations.  Moreover,
  6394 locale operations (renaming and type / term instantiation) are applied
  6395 to attribute arguments as expected.
  6396 
  6397 INCOMPATIBILITY of the ML interface: always pass Attrib.src instead of
  6398 actual attributes; rare situations may require Attrib.attribute to
  6399 embed those attributes into Attrib.src that lack concrete syntax.
  6400 Attribute implementations need to cooperate properly with the static
  6401 binding mechanism.  Basic parsers Args.XXX_typ/term/prop and
  6402 Attrib.XXX_thm etc. already do the right thing without further
  6403 intervention.  Only unusual applications -- such as "where" or "of"
  6404 (cf. src/Pure/Isar/attrib.ML), which process arguments depending both
  6405 on the context and the facts involved -- may have to assign parsed
  6406 values to argument tokens explicitly.
  6407 
  6408 * Changed parameter management in theorem generation for long goal
  6409 statements with 'includes'.  INCOMPATIBILITY: produces a different
  6410 theorem statement in rare situations.
  6411 
  6412 * Locale inspection command 'print_locale' omits notes elements.  Use
  6413 'print_locale!' to have them included in the output.
  6414 
  6415 
  6416 *** Provers ***
  6417 
  6418 * Provers/hypsubst.ML: improved version of the subst method, for
  6419 single-step rewriting: it now works in bound variable contexts. New is
  6420 'subst (asm)', for rewriting an assumption.  INCOMPATIBILITY: may
  6421 rewrite a different subterm than the original subst method, which is
  6422 still available as 'simplesubst'.
  6423 
  6424 * Provers/quasi.ML: new transitivity reasoners for transitivity only
  6425 and quasi orders.
  6426 
  6427 * Provers/trancl.ML: new transitivity reasoner for transitive and
  6428 reflexive-transitive closure of relations.
  6429 
  6430 * Provers/blast.ML: new reference depth_limit to make blast's depth
  6431 limit (previously hard-coded with a value of 20) user-definable.
  6432 
  6433 * Provers/simplifier.ML has been moved to Pure, where Simplifier.setup
  6434 is peformed already.  Object-logics merely need to finish their
  6435 initial simpset configuration as before.  INCOMPATIBILITY.
  6436 
  6437 
  6438 *** HOL ***
  6439 
  6440 * Symbolic syntax of Hilbert Choice Operator is now as follows:
  6441 
  6442   syntax (epsilon)
  6443     "_Eps" :: "[pttrn, bool] => 'a"    ("(3\<some>_./ _)" [0, 10] 10)
  6444 
  6445 The symbol \<some> is displayed as the alternative epsilon of LaTeX
  6446 and x-symbol; use option '-m epsilon' to get it actually printed.
  6447 Moreover, the mathematically important symbolic identifier \<epsilon>
  6448 becomes available as variable, constant etc.  INCOMPATIBILITY,
  6449 
  6450 * "x > y" abbreviates "y < x" and "x >= y" abbreviates "y <= x".
  6451 Similarly for all quantifiers: "ALL x > y" etc.  The x-symbol for >=
  6452 is \<ge>. New transitivity rules have been added to HOL/Orderings.thy to
  6453 support corresponding Isar calculations.
  6454 
  6455 * "{x:A. P}" abbreviates "{x. x:A & P}", and similarly for "\<in>"
  6456 instead of ":".
  6457 
  6458 * theory SetInterval: changed the syntax for open intervals:
  6459 
  6460   Old       New
  6461   {..n(}    {..<n}
  6462   {)n..}    {n<..}
  6463   {m..n(}   {m..<n}
  6464   {)m..n}   {m<..n}
  6465   {)m..n(}  {m<..<n}
  6466 
  6467 The old syntax is still supported but will disappear in the next
  6468 release.  For conversion use the following Emacs search and replace
  6469 patterns (these are not perfect but work quite well):
  6470 
  6471   {)\([^\.]*\)\.\.  ->  {\1<\.\.}
  6472   \.\.\([^(}]*\)(}  ->  \.\.<\1}
  6473 
  6474 * Theory Commutative_Ring (in Library): method comm_ring for proving
  6475 equalities in commutative rings; method 'algebra' provides a generic
  6476 interface.
  6477 
  6478 * Theory Finite_Set: changed the syntax for 'setsum', summation over
  6479 finite sets: "setsum (%x. e) A", which used to be "\<Sum>x:A. e", is
  6480 now either "SUM x:A. e" or "\<Sum>x \<in> A. e". The bound variable can
  6481 be a tuple pattern.
  6482 
  6483 Some new syntax forms are available:
  6484 
  6485   "\<Sum>x | P. e"      for     "setsum (%x. e) {x. P}"
  6486   "\<Sum>x = a..b. e"   for     "setsum (%x. e) {a..b}"
  6487   "\<Sum>x = a..<b. e"  for     "setsum (%x. e) {a..<b}"
  6488   "\<Sum>x < k. e"      for     "setsum (%x. e) {..<k}"
  6489 
  6490 The latter form "\<Sum>x < k. e" used to be based on a separate
  6491 function "Summation", which has been discontinued.
  6492 
  6493 * theory Finite_Set: in structured induction proofs, the insert case
  6494 is now 'case (insert x F)' instead of the old counterintuitive 'case
  6495 (insert F x)'.
  6496 
  6497 * The 'refute' command has been extended to support a much larger
  6498 fragment of HOL, including axiomatic type classes, constdefs and
  6499 typedefs, inductive datatypes and recursion.
  6500 
  6501 * New tactics 'sat' and 'satx' to prove propositional tautologies.
  6502 Requires zChaff with proof generation to be installed.  See
  6503 HOL/ex/SAT_Examples.thy for examples.
  6504 
  6505 * Datatype induction via method 'induct' now preserves the name of the
  6506 induction variable. For example, when proving P(xs::'a list) by
  6507 induction on xs, the induction step is now P(xs) ==> P(a#xs) rather
  6508 than P(list) ==> P(a#list) as previously.  Potential INCOMPATIBILITY
  6509 in unstructured proof scripts.
  6510 
  6511 * Reworked implementation of records.  Improved scalability for
  6512 records with many fields, avoiding performance problems for type
  6513 inference. Records are no longer composed of nested field types, but
  6514 of nested extension types. Therefore the record type only grows linear
  6515 in the number of extensions and not in the number of fields.  The
  6516 top-level (users) view on records is preserved.  Potential
  6517 INCOMPATIBILITY only in strange cases, where the theory depends on the
  6518 old record representation. The type generated for a record is called
  6519 <record_name>_ext_type.
  6520 
  6521 Flag record_quick_and_dirty_sensitive can be enabled to skip the
  6522 proofs triggered by a record definition or a simproc (if
  6523 quick_and_dirty is enabled).  Definitions of large records can take
  6524 quite long.
  6525 
  6526 New simproc record_upd_simproc for simplification of multiple record
  6527 updates enabled by default.  Moreover, trivial updates are also
  6528 removed: r(|x := x r|) = r.  INCOMPATIBILITY: old proofs break
  6529 occasionally, since simplification is more powerful by default.
  6530 
  6531 * typedef: proper support for polymorphic sets, which contain extra
  6532 type-variables in the term.
  6533 
  6534 * Simplifier: automatically reasons about transitivity chains
  6535 involving "trancl" (r^+) and "rtrancl" (r^*) by setting up tactics
  6536 provided by Provers/trancl.ML as additional solvers.  INCOMPATIBILITY:
  6537 old proofs break occasionally as simplification may now solve more
  6538 goals than previously.
  6539 
  6540 * Simplifier: converts x <= y into x = y if assumption y <= x is
  6541 present.  Works for all partial orders (class "order"), in particular
  6542 numbers and sets.  For linear orders (e.g. numbers) it treats ~ x < y
  6543 just like y <= x.
  6544 
  6545 * Simplifier: new simproc for "let x = a in f x".  If a is a free or
  6546 bound variable or a constant then the let is unfolded.  Otherwise
  6547 first a is simplified to b, and then f b is simplified to g. If
  6548 possible we abstract b from g arriving at "let x = b in h x",
  6549 otherwise we unfold the let and arrive at g.  The simproc can be
  6550 enabled/disabled by the reference use_let_simproc.  Potential
  6551 INCOMPATIBILITY since simplification is more powerful by default.
  6552 
  6553 * Classical reasoning: the meson method now accepts theorems as arguments.
  6554 
  6555 * Prover support: pre-release of the Isabelle-ATP linkup, which runs background
  6556 jobs to provide advice on the provability of subgoals.
  6557 
  6558 * Theory OrderedGroup and Ring_and_Field: various additions and
  6559 improvements to faciliate calculations involving equalities and
  6560 inequalities.
  6561 
  6562 The following theorems have been eliminated or modified
  6563 (INCOMPATIBILITY):
  6564 
  6565   abs_eq             now named abs_of_nonneg
  6566   abs_of_ge_0        now named abs_of_nonneg
  6567   abs_minus_eq       now named abs_of_nonpos
  6568   imp_abs_id         now named abs_of_nonneg
  6569   imp_abs_neg_id     now named abs_of_nonpos
  6570   mult_pos           now named mult_pos_pos
  6571   mult_pos_le        now named mult_nonneg_nonneg
  6572   mult_pos_neg_le    now named mult_nonneg_nonpos
  6573   mult_pos_neg2_le   now named mult_nonneg_nonpos2
  6574   mult_neg           now named mult_neg_neg
  6575   mult_neg_le        now named mult_nonpos_nonpos
  6576 
  6577 * The following lemmas in Ring_and_Field have been added to the simplifier:
  6578      
  6579      zero_le_square
  6580      not_square_less_zero 
  6581 
  6582   The following lemmas have been deleted from Real/RealPow:
  6583   
  6584      realpow_zero_zero
  6585      realpow_two
  6586      realpow_less
  6587      zero_le_power
  6588      realpow_two_le
  6589      abs_realpow_two
  6590      realpow_two_abs     
  6591 
  6592 * Theory Parity: added rules for simplifying exponents.
  6593 
  6594 * Theory List:
  6595 
  6596 The following theorems have been eliminated or modified
  6597 (INCOMPATIBILITY):
  6598 
  6599   list_all_Nil       now named list_all.simps(1)
  6600   list_all_Cons      now named list_all.simps(2)
  6601   list_all_conv      now named list_all_iff
  6602   set_mem_eq         now named mem_iff
  6603 
  6604 * Theories SetsAndFunctions and BigO (see HOL/Library) support
  6605 asymptotic "big O" calculations.  See the notes in BigO.thy.
  6606 
  6607 
  6608 *** HOL-Complex ***
  6609 
  6610 * Theory RealDef: better support for embedding natural numbers and
  6611 integers in the reals.
  6612 
  6613 The following theorems have been eliminated or modified
  6614 (INCOMPATIBILITY):
  6615 
  6616   exp_ge_add_one_self  now requires no hypotheses
  6617   real_of_int_add      reversed direction of equality (use [symmetric])
  6618   real_of_int_minus    reversed direction of equality (use [symmetric])
  6619   real_of_int_diff     reversed direction of equality (use [symmetric])
  6620   real_of_int_mult     reversed direction of equality (use [symmetric])
  6621 
  6622 * Theory RComplete: expanded support for floor and ceiling functions.
  6623 
  6624 * Theory Ln is new, with properties of the natural logarithm
  6625 
  6626 * Hyperreal: There is a new type constructor "star" for making
  6627 nonstandard types.  The old type names are now type synonyms:
  6628 
  6629   hypreal = real star
  6630   hypnat = nat star
  6631   hcomplex = complex star
  6632 
  6633 * Hyperreal: Many groups of similarly-defined constants have been
  6634 replaced by polymorphic versions (INCOMPATIBILITY):
  6635 
  6636   star_of <-- hypreal_of_real, hypnat_of_nat, hcomplex_of_complex
  6637 
  6638   starset      <-- starsetNat, starsetC
  6639   *s*          <-- *sNat*, *sc*
  6640   starset_n    <-- starsetNat_n, starsetC_n
  6641   *sn*         <-- *sNatn*, *scn*
  6642   InternalSets <-- InternalNatSets, InternalCSets
  6643 
  6644   starfun      <-- starfun{Nat,Nat2,C,RC,CR}
  6645   *f*          <-- *fNat*, *fNat2*, *fc*, *fRc*, *fcR*
  6646   starfun_n    <-- starfun{Nat,Nat2,C,RC,CR}_n
  6647   *fn*         <-- *fNatn*, *fNat2n*, *fcn*, *fRcn*, *fcRn*
  6648   InternalFuns <-- InternalNatFuns, InternalNatFuns2, Internal{C,RC,CR}Funs
  6649 
  6650 * Hyperreal: Many type-specific theorems have been removed in favor of
  6651 theorems specific to various axiomatic type classes (INCOMPATIBILITY):
  6652 
  6653   add_commute <-- {hypreal,hypnat,hcomplex}_add_commute
  6654   add_assoc   <-- {hypreal,hypnat,hcomplex}_add_assocs
  6655   OrderedGroup.add_0 <-- {hypreal,hypnat,hcomplex}_add_zero_left
  6656   OrderedGroup.add_0_right <-- {hypreal,hcomplex}_add_zero_right
  6657   right_minus <-- hypreal_add_minus
  6658   left_minus <-- {hypreal,hcomplex}_add_minus_left
  6659   mult_commute <-- {hypreal,hypnat,hcomplex}_mult_commute
  6660   mult_assoc <-- {hypreal,hypnat,hcomplex}_mult_assoc
  6661   mult_1_left <-- {hypreal,hypnat}_mult_1, hcomplex_mult_one_left
  6662   mult_1_right <-- hcomplex_mult_one_right
  6663   mult_zero_left <-- hcomplex_mult_zero_left
  6664   left_distrib <-- {hypreal,hypnat,hcomplex}_add_mult_distrib
  6665   right_distrib <-- hypnat_add_mult_distrib2
  6666   zero_neq_one <-- {hypreal,hypnat,hcomplex}_zero_not_eq_one
  6667   right_inverse <-- hypreal_mult_inverse
  6668   left_inverse <-- hypreal_mult_inverse_left, hcomplex_mult_inv_left
  6669   order_refl <-- {hypreal,hypnat}_le_refl
  6670   order_trans <-- {hypreal,hypnat}_le_trans
  6671   order_antisym <-- {hypreal,hypnat}_le_anti_sym
  6672   order_less_le <-- {hypreal,hypnat}_less_le
  6673   linorder_linear <-- {hypreal,hypnat}_le_linear
  6674   add_left_mono <-- {hypreal,hypnat}_add_left_mono
  6675   mult_strict_left_mono <-- {hypreal,hypnat}_mult_less_mono2
  6676   add_nonneg_nonneg <-- hypreal_le_add_order
  6677 
  6678 * Hyperreal: Separate theorems having to do with type-specific
  6679 versions of constants have been merged into theorems that apply to the
  6680 new polymorphic constants (INCOMPATIBILITY):
  6681 
  6682   STAR_UNIV_set <-- {STAR_real,NatStar_real,STARC_complex}_set
  6683   STAR_empty_set <-- {STAR,NatStar,STARC}_empty_set
  6684   STAR_Un <-- {STAR,NatStar,STARC}_Un
  6685   STAR_Int <-- {STAR,NatStar,STARC}_Int
  6686   STAR_Compl <-- {STAR,NatStar,STARC}_Compl
  6687   STAR_subset <-- {STAR,NatStar,STARC}_subset
  6688   STAR_mem <-- {STAR,NatStar,STARC}_mem
  6689   STAR_mem_Compl <-- {STAR,STARC}_mem_Compl
  6690   STAR_diff <-- {STAR,STARC}_diff
  6691   STAR_star_of_image_subset <-- {STAR_hypreal_of_real, NatStar_hypreal_of_real,
  6692     STARC_hcomplex_of_complex}_image_subset
  6693   starset_n_Un <-- starset{Nat,C}_n_Un
  6694   starset_n_Int <-- starset{Nat,C}_n_Int
  6695   starset_n_Compl <-- starset{Nat,C}_n_Compl
  6696   starset_n_diff <-- starset{Nat,C}_n_diff
  6697   InternalSets_Un <-- Internal{Nat,C}Sets_Un
  6698   InternalSets_Int <-- Internal{Nat,C}Sets_Int
  6699   InternalSets_Compl <-- Internal{Nat,C}Sets_Compl
  6700   InternalSets_diff <-- Internal{Nat,C}Sets_diff
  6701   InternalSets_UNIV_diff <-- Internal{Nat,C}Sets_UNIV_diff
  6702   InternalSets_starset_n <-- Internal{Nat,C}Sets_starset{Nat,C}_n
  6703   starset_starset_n_eq <-- starset{Nat,C}_starset{Nat,C}_n_eq
  6704   starset_n_starset <-- starset{Nat,C}_n_starset{Nat,C}
  6705   starfun_n_starfun <-- starfun{Nat,Nat2,C,RC,CR}_n_starfun{Nat,Nat2,C,RC,CR}
  6706   starfun <-- starfun{Nat,Nat2,C,RC,CR}
  6707   starfun_mult <-- starfun{Nat,Nat2,C,RC,CR}_mult
  6708   starfun_add <-- starfun{Nat,Nat2,C,RC,CR}_add
  6709   starfun_minus <-- starfun{Nat,Nat2,C,RC,CR}_minus
  6710   starfun_diff <-- starfun{C,RC,CR}_diff
  6711   starfun_o <-- starfun{NatNat2,Nat2,_stafunNat,C,C_starfunRC,_starfunCR}_o
  6712   starfun_o2 <-- starfun{NatNat2,_stafunNat,C,C_starfunRC,_starfunCR}_o2
  6713   starfun_const_fun <-- starfun{Nat,Nat2,C,RC,CR}_const_fun
  6714   starfun_inverse <-- starfun{Nat,C,RC,CR}_inverse
  6715   starfun_eq <-- starfun{Nat,Nat2,C,RC,CR}_eq
  6716   starfun_eq_iff <-- starfun{C,RC,CR}_eq_iff
  6717   starfun_Id <-- starfunC_Id
  6718   starfun_approx <-- starfun{Nat,CR}_approx
  6719   starfun_capprox <-- starfun{C,RC}_capprox
  6720   starfun_abs <-- starfunNat_rabs
  6721   starfun_lambda_cancel <-- starfun{C,CR,RC}_lambda_cancel
  6722   starfun_lambda_cancel2 <-- starfun{C,CR,RC}_lambda_cancel2
  6723   starfun_mult_HFinite_approx <-- starfunCR_mult_HFinite_capprox
  6724   starfun_mult_CFinite_capprox <-- starfun{C,RC}_mult_CFinite_capprox
  6725   starfun_add_capprox <-- starfun{C,RC}_add_capprox
  6726   starfun_add_approx <-- starfunCR_add_approx
  6727   starfun_inverse_inverse <-- starfunC_inverse_inverse
  6728   starfun_divide <-- starfun{C,CR,RC}_divide
  6729   starfun_n <-- starfun{Nat,C}_n
  6730   starfun_n_mult <-- starfun{Nat,C}_n_mult
  6731   starfun_n_add <-- starfun{Nat,C}_n_add
  6732   starfun_n_add_minus <-- starfunNat_n_add_minus
  6733   starfun_n_const_fun <-- starfun{Nat,C}_n_const_fun
  6734   starfun_n_minus <-- starfun{Nat,C}_n_minus
  6735   starfun_n_eq <-- starfun{Nat,C}_n_eq
  6736 
  6737   star_n_add <-- {hypreal,hypnat,hcomplex}_add
  6738   star_n_minus <-- {hypreal,hcomplex}_minus
  6739   star_n_diff <-- {hypreal,hcomplex}_diff
  6740   star_n_mult <-- {hypreal,hcomplex}_mult
  6741   star_n_inverse <-- {hypreal,hcomplex}_inverse
  6742   star_n_le <-- {hypreal,hypnat}_le
  6743   star_n_less <-- {hypreal,hypnat}_less
  6744   star_n_zero_num <-- {hypreal,hypnat,hcomplex}_zero_num
  6745   star_n_one_num <-- {hypreal,hypnat,hcomplex}_one_num
  6746   star_n_abs <-- hypreal_hrabs
  6747   star_n_divide <-- hcomplex_divide
  6748 
  6749   star_of_add <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_add
  6750   star_of_minus <-- {hypreal_of_real,hcomplex_of_complex}_minus
  6751   star_of_diff <-- hypreal_of_real_diff
  6752   star_of_mult <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_mult
  6753   star_of_one <-- {hypreal_of_real,hcomplex_of_complex}_one
  6754   star_of_zero <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_zero
  6755   star_of_le <-- {hypreal_of_real,hypnat_of_nat}_le_iff
  6756   star_of_less <-- {hypreal_of_real,hypnat_of_nat}_less_iff
  6757   star_of_eq <-- {hypreal_of_real,hypnat_of_nat,hcomplex_of_complex}_eq_iff
  6758   star_of_inverse <-- {hypreal_of_real,hcomplex_of_complex}_inverse
  6759   star_of_divide <-- {hypreal_of_real,hcomplex_of_complex}_divide
  6760   star_of_of_nat <-- {hypreal_of_real,hcomplex_of_complex}_of_nat
  6761   star_of_of_int <-- {hypreal_of_real,hcomplex_of_complex}_of_int
  6762   star_of_number_of <-- {hypreal,hcomplex}_number_of
  6763   star_of_number_less <-- number_of_less_hypreal_of_real_iff
  6764   star_of_number_le <-- number_of_le_hypreal_of_real_iff
  6765   star_of_eq_number <-- hypreal_of_real_eq_number_of_iff
  6766   star_of_less_number <-- hypreal_of_real_less_number_of_iff
  6767   star_of_le_number <-- hypreal_of_real_le_number_of_iff
  6768   star_of_power <-- hypreal_of_real_power
  6769   star_of_eq_0 <-- hcomplex_of_complex_zero_iff
  6770 
  6771 * Hyperreal: new method "transfer" that implements the transfer
  6772 principle of nonstandard analysis. With a subgoal that mentions
  6773 nonstandard types like "'a star", the command "apply transfer"
  6774 replaces it with an equivalent one that mentions only standard types.
  6775 To be successful, all free variables must have standard types; non-
  6776 standard variables must have explicit universal quantifiers.
  6777 
  6778 * Hyperreal: A theory of Taylor series.
  6779 
  6780 
  6781 *** HOLCF ***
  6782 
  6783 * Discontinued special version of 'constdefs' (which used to support
  6784 continuous functions) in favor of the general Pure one with full
  6785 type-inference.
  6786 
  6787 * New simplification procedure for solving continuity conditions; it
  6788 is much faster on terms with many nested lambda abstractions (cubic
  6789 instead of exponential time).
  6790 
  6791 * New syntax for domain package: selector names are now optional.
  6792 Parentheses should be omitted unless argument is lazy, for example:
  6793 
  6794   domain 'a stream = cons "'a" (lazy "'a stream")
  6795 
  6796 * New command 'fixrec' for defining recursive functions with pattern
  6797 matching; defining multiple functions with mutual recursion is also
  6798 supported.  Patterns may include the constants cpair, spair, up, sinl,
  6799 sinr, or any data constructor defined by the domain package. The given
  6800 equations are proven as rewrite rules. See HOLCF/ex/Fixrec_ex.thy for
  6801 syntax and examples.
  6802 
  6803 * New commands 'cpodef' and 'pcpodef' for defining predicate subtypes
  6804 of cpo and pcpo types. Syntax is exactly like the 'typedef' command,
  6805 but the proof obligation additionally includes an admissibility
  6806 requirement. The packages generate instances of class cpo or pcpo,
  6807 with continuity and strictness theorems for Rep and Abs.
  6808 
  6809 * HOLCF: Many theorems have been renamed according to a more standard naming
  6810 scheme (INCOMPATIBILITY):
  6811 
  6812   foo_inject:  "foo$x = foo$y ==> x = y"
  6813   foo_eq:      "(foo$x = foo$y) = (x = y)"
  6814   foo_less:    "(foo$x << foo$y) = (x << y)"
  6815   foo_strict:  "foo$UU = UU"
  6816   foo_defined: "... ==> foo$x ~= UU"
  6817   foo_defined_iff: "(foo$x = UU) = (x = UU)"
  6818 
  6819 
  6820 *** ZF ***
  6821 
  6822 * ZF/ex: theories Group and Ring provide examples in abstract algebra,
  6823 including the First Isomorphism Theorem (on quotienting by the kernel
  6824 of a homomorphism).
  6825 
  6826 * ZF/Simplifier: install second copy of type solver that actually
  6827 makes use of TC rules declared to Isar proof contexts (or locales);
  6828 the old version is still required for ML proof scripts.
  6829 
  6830 
  6831 *** Cube ***
  6832 
  6833 * Converted to Isar theory format; use locales instead of axiomatic
  6834 theories.
  6835 
  6836 
  6837 *** ML ***
  6838 
  6839 * Pure/library.ML: added ##>, ##>>, #>> -- higher-order counterparts
  6840 for ||>, ||>>, |>>,
  6841 
  6842 * Pure/library.ML no longer defines its own option datatype, but uses
  6843 that of the SML basis, which has constructors NONE and SOME instead of
  6844 None and Some, as well as exception Option.Option instead of OPTION.
  6845 The functions the, if_none, is_some, is_none have been adapted
  6846 accordingly, while Option.map replaces apsome.
  6847 
  6848 * Pure/library.ML: the exception LIST has been given up in favour of
  6849 the standard exceptions Empty and Subscript, as well as
  6850 Library.UnequalLengths.  Function like Library.hd and Library.tl are
  6851 superceded by the standard hd and tl functions etc.
  6852 
  6853 A number of basic list functions are no longer exported to the ML
  6854 toplevel, as they are variants of predefined functions.  The following
  6855 suggests how one can translate existing code:
  6856 
  6857     rev_append xs ys = List.revAppend (xs, ys)
  6858     nth_elem (i, xs) = List.nth (xs, i)
  6859     last_elem xs = List.last xs
  6860     flat xss = List.concat xss
  6861     seq fs = List.app fs
  6862     partition P xs = List.partition P xs
  6863     mapfilter f xs = List.mapPartial f xs
  6864 
  6865 * Pure/library.ML: several combinators for linear functional
  6866 transformations, notably reverse application and composition:
  6867 
  6868   x |> f                f #> g
  6869   (x, y) |-> f          f #-> g
  6870 
  6871 * Pure/library.ML: introduced/changed precedence of infix operators:
  6872 
  6873   infix 1 |> |-> ||> ||>> |>> |>>> #> #->;
  6874   infix 2 ?;
  6875   infix 3 o oo ooo oooo;
  6876   infix 4 ~~ upto downto;
  6877 
  6878 Maybe INCOMPATIBILITY when any of those is used in conjunction with other
  6879 infix operators.
  6880 
  6881 * Pure/library.ML: natural list combinators fold, fold_rev, and
  6882 fold_map support linear functional transformations and nesting.  For
  6883 example:
  6884 
  6885   fold f [x1, ..., xN] y =
  6886     y |> f x1 |> ... |> f xN
  6887 
  6888   (fold o fold) f [xs1, ..., xsN] y =
  6889     y |> fold f xs1 |> ... |> fold f xsN
  6890 
  6891   fold f [x1, ..., xN] =
  6892     f x1 #> ... #> f xN
  6893 
  6894   (fold o fold) f [xs1, ..., xsN] =
  6895     fold f xs1 #> ... #> fold f xsN
  6896 
  6897 * Pure/library.ML: the following selectors on type 'a option are
  6898 available:
  6899 
  6900   the:               'a option -> 'a  (*partial*)
  6901   these:             'a option -> 'a  where 'a = 'b list
  6902   the_default: 'a -> 'a option -> 'a
  6903   the_list:          'a option -> 'a list
  6904 
  6905 * Pure/General: structure AList (cf. Pure/General/alist.ML) provides
  6906 basic operations for association lists, following natural argument
  6907 order; moreover the explicit equality predicate passed here avoids
  6908 potentially expensive polymorphic runtime equality checks.
  6909 The old functions may be expressed as follows:
  6910 
  6911   assoc = uncurry (AList.lookup (op =))
  6912   assocs = these oo AList.lookup (op =)
  6913   overwrite = uncurry (AList.update (op =)) o swap
  6914 
  6915 * Pure/General: structure AList (cf. Pure/General/alist.ML) provides
  6916 
  6917   val make: ('a -> 'b) -> 'a list -> ('a * 'b) list
  6918   val find: ('a * 'b -> bool) -> ('c * 'b) list -> 'a -> 'c list
  6919 
  6920 replacing make_keylist and keyfilter (occassionally used)
  6921 Naive rewrites:
  6922 
  6923   make_keylist = AList.make
  6924   keyfilter = AList.find (op =)
  6925 
  6926 * eq_fst and eq_snd now take explicit equality parameter, thus
  6927   avoiding eqtypes. Naive rewrites:
  6928 
  6929     eq_fst = eq_fst (op =)
  6930     eq_snd = eq_snd (op =)
  6931 
  6932 * Removed deprecated apl and apr (rarely used).
  6933   Naive rewrites:
  6934 
  6935     apl (n, op) =>>= curry op n
  6936     apr (op, m) =>>= fn n => op (n, m)
  6937 
  6938 * Pure/General: structure OrdList (cf. Pure/General/ord_list.ML)
  6939 provides a reasonably efficient light-weight implementation of sets as
  6940 lists.
  6941 
  6942 * Pure/General: generic tables (cf. Pure/General/table.ML) provide a
  6943 few new operations; existing lookup and update are now curried to
  6944 follow natural argument order (for use with fold etc.);
  6945 INCOMPATIBILITY, use (uncurry Symtab.lookup) etc. as last resort.
  6946 
  6947 * Pure/General: output via the Isabelle channels of
  6948 writeln/warning/error etc. is now passed through Output.output, with a
  6949 hook for arbitrary transformations depending on the print_mode
  6950 (cf. Output.add_mode -- the first active mode that provides a output
  6951 function wins).  Already formatted output may be embedded into further
  6952 text via Output.raw; the result of Pretty.string_of/str_of and derived
  6953 functions (string_of_term/cterm/thm etc.) is already marked raw to
  6954 accommodate easy composition of diagnostic messages etc.  Programmers
  6955 rarely need to care about Output.output or Output.raw at all, with
  6956 some notable exceptions: Output.output is required when bypassing the
  6957 standard channels (writeln etc.), or in token translations to produce
  6958 properly formatted results; Output.raw is required when capturing
  6959 already output material that will eventually be presented to the user
  6960 a second time.  For the default print mode, both Output.output and
  6961 Output.raw have no effect.
  6962 
  6963 * Pure/General: Output.time_accumulator NAME creates an operator ('a
  6964 -> 'b) -> 'a -> 'b to measure runtime and count invocations; the
  6965 cumulative results are displayed at the end of a batch session.
  6966 
  6967 * Pure/General: File.sysify_path and File.quote_sysify path have been
  6968 replaced by File.platform_path and File.shell_path (with appropriate
  6969 hooks).  This provides a clean interface for unusual systems where the
  6970 internal and external process view of file names are different.
  6971 
  6972 * Pure: more efficient orders for basic syntactic entities: added
  6973 fast_string_ord, fast_indexname_ord, fast_term_ord; changed sort_ord
  6974 and typ_ord to use fast_string_ord and fast_indexname_ord (term_ord is
  6975 NOT affected); structures Symtab, Vartab, Typtab, Termtab use the fast
  6976 orders now -- potential INCOMPATIBILITY for code that depends on a
  6977 particular order for Symtab.keys, Symtab.dest, etc. (consider using
  6978 Library.sort_strings on result).
  6979 
  6980 * Pure/term.ML: combinators fold_atyps, fold_aterms, fold_term_types,
  6981 fold_types traverse types/terms from left to right, observing natural
  6982 argument order.  Supercedes previous foldl_XXX versions, add_frees,
  6983 add_vars etc. have been adapted as well: INCOMPATIBILITY.
  6984 
  6985 * Pure: name spaces have been refined, with significant changes of the
  6986 internal interfaces -- INCOMPATIBILITY.  Renamed cond_extern(_table)
  6987 to extern(_table).  The plain name entry path is superceded by a
  6988 general 'naming' context, which also includes the 'policy' to produce
  6989 a fully qualified name and external accesses of a fully qualified
  6990 name; NameSpace.extend is superceded by context dependent
  6991 Sign.declare_name.  Several theory and proof context operations modify
  6992 the naming context.  Especially note Theory.restore_naming and
  6993 ProofContext.restore_naming to get back to a sane state; note that
  6994 Theory.add_path is no longer sufficient to recover from
  6995 Theory.absolute_path in particular.
  6996 
  6997 * Pure: new flags short_names (default false) and unique_names
  6998 (default true) for controlling output of qualified names.  If
  6999 short_names is set, names are printed unqualified.  If unique_names is
  7000 reset, the name prefix is reduced to the minimum required to achieve
  7001 the original result when interning again, even if there is an overlap
  7002 with earlier declarations.
  7003 
  7004 * Pure/TheoryDataFun: change of the argument structure; 'prep_ext' is
  7005 now 'extend', and 'merge' gets an additional Pretty.pp argument
  7006 (useful for printing error messages).  INCOMPATIBILITY.
  7007 
  7008 * Pure: major reorganization of the theory context.  Type Sign.sg and
  7009 Theory.theory are now identified, referring to the universal
  7010 Context.theory (see Pure/context.ML).  Actual signature and theory
  7011 content is managed as theory data.  The old code and interfaces were
  7012 spread over many files and structures; the new arrangement introduces
  7013 considerable INCOMPATIBILITY to gain more clarity:
  7014 
  7015   Context -- theory management operations (name, identity, inclusion,
  7016     parents, ancestors, merge, etc.), plus generic theory data;
  7017 
  7018   Sign -- logical signature and syntax operations (declaring consts,
  7019     types, etc.), plus certify/read for common entities;
  7020 
  7021   Theory -- logical theory operations (stating axioms, definitions,
  7022     oracles), plus a copy of logical signature operations (consts,
  7023     types, etc.); also a few basic management operations (Theory.copy,
  7024     Theory.merge, etc.)
  7025 
  7026 The most basic sign_of operations (Theory.sign_of, Thm.sign_of_thm
  7027 etc.) as well as the sign field in Thm.rep_thm etc. have been retained
  7028 for convenience -- they merely return the theory.
  7029 
  7030 * Pure: type Type.tsig is superceded by theory in most interfaces.
  7031 
  7032 * Pure: the Isar proof context type is already defined early in Pure
  7033 as Context.proof (note that ProofContext.context and Proof.context are
  7034 aliases, where the latter is the preferred name).  This enables other
  7035 Isabelle components to refer to that type even before Isar is present.
  7036 
  7037 * Pure/sign/theory: discontinued named name spaces (i.e. classK,
  7038 typeK, constK, axiomK, oracleK), but provide explicit operations for
  7039 any of these kinds.  For example, Sign.intern typeK is now
  7040 Sign.intern_type, Theory.hide_space Sign.typeK is now
  7041 Theory.hide_types.  Also note that former
  7042 Theory.hide_classes/types/consts are now
  7043 Theory.hide_classes_i/types_i/consts_i, while the non '_i' versions
  7044 internalize their arguments!  INCOMPATIBILITY.
  7045 
  7046 * Pure: get_thm interface (of PureThy and ProofContext) expects
  7047 datatype thmref (with constructors Name and NameSelection) instead of
  7048 plain string -- INCOMPATIBILITY;
  7049 
  7050 * Pure: cases produced by proof methods specify options, where NONE
  7051 means to remove case bindings -- INCOMPATIBILITY in
  7052 (RAW_)METHOD_CASES.
  7053 
  7054 * Pure: the following operations retrieve axioms or theorems from a
  7055 theory node or theory hierarchy, respectively:
  7056 
  7057   Theory.axioms_of: theory -> (string * term) list
  7058   Theory.all_axioms_of: theory -> (string * term) list
  7059   PureThy.thms_of: theory -> (string * thm) list
  7060   PureThy.all_thms_of: theory -> (string * thm) list
  7061 
  7062 * Pure: print_tac now outputs the goal through the trace channel.
  7063 
  7064 * Isar toplevel: improved diagnostics, mostly for Poly/ML only.
  7065 Reference Toplevel.debug (default false) controls detailed printing
  7066 and tracing of low-level exceptions; Toplevel.profiling (default 0)
  7067 controls execution profiling -- set to 1 for time and 2 for space
  7068 (both increase the runtime).
  7069 
  7070 * Isar session: The initial use of ROOT.ML is now always timed,
  7071 i.e. the log will show the actual process times, in contrast to the
  7072 elapsed wall-clock time that the outer shell wrapper produces.
  7073 
  7074 * Simplifier: improved handling of bound variables (nameless
  7075 representation, avoid allocating new strings).  Simprocs that invoke
  7076 the Simplifier recursively should use Simplifier.inherit_bounds to
  7077 avoid local name clashes.  Failure to do so produces warnings
  7078 "Simplifier: renamed bound variable ..."; set Simplifier.debug_bounds
  7079 for further details.
  7080 
  7081 * ML functions legacy_bindings and use_legacy_bindings produce ML fact
  7082 bindings for all theorems stored within a given theory; this may help
  7083 in porting non-Isar theories to Isar ones, while keeping ML proof
  7084 scripts for the time being.
  7085 
  7086 * ML operator HTML.with_charset specifies the charset begin used for
  7087 generated HTML files.  For example:
  7088 
  7089   HTML.with_charset "utf-8" use_thy "Hebrew";
  7090   HTML.with_charset "utf-8" use_thy "Chinese";
  7091 
  7092 
  7093 *** System ***
  7094 
  7095 * Allow symlinks to all proper Isabelle executables (Isabelle,
  7096 isabelle, isatool etc.).
  7097 
  7098 * ISABELLE_DOC_FORMAT setting specifies preferred document format (for
  7099 isatool doc, isatool mkdir, display_drafts etc.).
  7100 
  7101 * isatool usedir: option -f allows specification of the ML file to be
  7102 used by Isabelle; default is ROOT.ML.
  7103 
  7104 * New isatool version outputs the version identifier of the Isabelle
  7105 distribution being used.
  7106 
  7107 * HOL: new isatool dimacs2hol converts files in DIMACS CNF format
  7108 (containing Boolean satisfiability problems) into Isabelle/HOL
  7109 theories.
  7110 
  7111 
  7112 
  7113 New in Isabelle2004 (April 2004)
  7114 --------------------------------
  7115 
  7116 *** General ***
  7117 
  7118 * Provers/order.ML:  new efficient reasoner for partial and linear orders.
  7119   Replaces linorder.ML.
  7120 
  7121 * Pure: Greek letters (except small lambda, \<lambda>), as well as Gothic
  7122   (\<aa>...\<zz>\<AA>...\<ZZ>), calligraphic (\<A>...\<Z>), and Euler
  7123   (\<a>...\<z>), are now considered normal letters, and can therefore
  7124   be used anywhere where an ASCII letter (a...zA...Z) has until
  7125   now. COMPATIBILITY: This obviously changes the parsing of some
  7126   terms, especially where a symbol has been used as a binder, say
  7127   '\<Pi>x. ...', which is now a type error since \<Pi>x will be parsed
  7128   as an identifier.  Fix it by inserting a space around former
  7129   symbols.  Call 'isatool fixgreek' to try to fix parsing errors in
  7130   existing theory and ML files.
  7131 
  7132 * Pure: Macintosh and Windows line-breaks are now allowed in theory files.
  7133 
  7134 * Pure: single letter sub/superscripts (\<^isub> and \<^isup>) are now
  7135   allowed in identifiers. Similar to Greek letters \<^isub> is now considered
  7136   a normal (but invisible) letter. For multiple letter subscripts repeat
  7137   \<^isub> like this: x\<^isub>1\<^isub>2.
  7138 
  7139 * Pure: There are now sub-/superscripts that can span more than one
  7140   character. Text between \<^bsub> and \<^esub> is set in subscript in
  7141   ProofGeneral and LaTeX, text between \<^bsup> and \<^esup> in
  7142   superscript. The new control characters are not identifier parts.
  7143 
  7144 * Pure: Control-symbols of the form \<^raw:...> will literally print the
  7145   content of "..." to the latex file instead of \isacntrl... . The "..."
  7146   may consist of any printable characters excluding the end bracket >.
  7147 
  7148 * Pure: Using new Isar command "finalconsts" (or the ML functions
  7149   Theory.add_finals or Theory.add_finals_i) it is now possible to
  7150   declare constants "final", which prevents their being given a definition
  7151   later.  It is useful for constants whose behaviour is fixed axiomatically
  7152   rather than definitionally, such as the meta-logic connectives.
  7153 
  7154 * Pure: 'instance' now handles general arities with general sorts
  7155   (i.e. intersections of classes),
  7156 
  7157 * Presentation: generated HTML now uses a CSS style sheet to make layout
  7158   (somewhat) independent of content. It is copied from lib/html/isabelle.css.
  7159   It can be changed to alter the colors/layout of generated pages.
  7160 
  7161 
  7162 *** Isar ***
  7163 
  7164 * Tactic emulation methods rule_tac, erule_tac, drule_tac, frule_tac,
  7165   cut_tac, subgoal_tac and thin_tac:
  7166   - Now understand static (Isar) contexts.  As a consequence, users of Isar
  7167     locales are no longer forced to write Isar proof scripts.
  7168     For details see Isar Reference Manual, paragraph 4.3.2: Further tactic
  7169     emulations.
  7170   - INCOMPATIBILITY: names of variables to be instantiated may no
  7171     longer be enclosed in quotes.  Instead, precede variable name with `?'.
  7172     This is consistent with the instantiation attribute "where".
  7173 
  7174 * Attributes "where" and "of":
  7175   - Now take type variables of instantiated theorem into account when reading
  7176     the instantiation string.  This fixes a bug that caused instantiated
  7177     theorems to have too special types in some circumstances.
  7178   - "where" permits explicit instantiations of type variables.
  7179 
  7180 * Calculation commands "moreover" and "also" no longer interfere with
  7181   current facts ("this"), admitting arbitrary combinations with "then"
  7182   and derived forms.
  7183 
  7184 * Locales:
  7185   - Goal statements involving the context element "includes" no longer
  7186     generate theorems with internal delta predicates (those ending on
  7187     "_axioms") in the premise.
  7188     Resolve particular premise with <locale>.intro to obtain old form.
  7189   - Fixed bug in type inference ("unify_frozen") that prevented mix of target
  7190     specification and "includes" elements in goal statement.
  7191   - Rule sets <locale>.intro and <locale>.axioms no longer declared as
  7192     [intro?] and [elim?] (respectively) by default.
  7193   - Experimental command for instantiation of locales in proof contexts:
  7194         instantiate <label>[<attrs>]: <loc>
  7195     Instantiates locale <loc> and adds all its theorems to the current context
  7196     taking into account their attributes.  Label and attrs are optional
  7197     modifiers, like in theorem declarations.  If present, names of
  7198     instantiated theorems are qualified with <label>, and the attributes
  7199     <attrs> are applied after any attributes these theorems might have already.
  7200       If the locale has assumptions, a chained fact of the form
  7201     "<loc> t1 ... tn" is expected from which instantiations of the parameters
  7202     are derived.  The command does not support old-style locales declared
  7203     with "locale (open)".
  7204       A few (very simple) examples can be found in FOL/ex/LocaleInst.thy.
  7205 
  7206 * HOL: Tactic emulation methods induct_tac and case_tac understand static
  7207   (Isar) contexts.
  7208 
  7209 
  7210 *** HOL ***
  7211 
  7212 * Proof import: new image HOL4 contains the imported library from
  7213   the HOL4 system with about 2500 theorems. It is imported by
  7214   replaying proof terms produced by HOL4 in Isabelle. The HOL4 image
  7215   can be used like any other Isabelle image.  See
  7216   HOL/Import/HOL/README for more information.
  7217 
  7218 * Simplifier:
  7219   - Much improved handling of linear and partial orders.
  7220     Reasoners for linear and partial orders are set up for type classes
  7221     "linorder" and "order" respectively, and are added to the default simpset
  7222     as solvers.  This means that the simplifier can build transitivity chains
  7223     to solve goals from the assumptions.
  7224   - INCOMPATIBILITY: old proofs break occasionally.  Typically, applications
  7225     of blast or auto after simplification become unnecessary because the goal
  7226     is solved by simplification already.
  7227 
  7228 * Numerics: new theory Ring_and_Field contains over 250 basic numerical laws,
  7229     all proved in axiomatic type classes for semirings, rings and fields.
  7230 
  7231 * Numerics:
  7232   - Numeric types (nat, int, and in HOL-Complex rat, real, complex, etc.) are
  7233     now formalized using the Ring_and_Field theory mentioned above.
  7234   - INCOMPATIBILITY: simplification and arithmetic behaves somewhat differently
  7235     than before, because now they are set up once in a generic manner.
  7236   - INCOMPATIBILITY: many type-specific arithmetic laws have gone.
  7237     Look for the general versions in Ring_and_Field (and Power if they concern
  7238     exponentiation).
  7239 
  7240 * Type "rat" of the rational numbers is now available in HOL-Complex.
  7241 
  7242 * Records:
  7243   - Record types are now by default printed with their type abbreviation
  7244     instead of the list of all field types. This can be configured via
  7245     the reference "print_record_type_abbr".
  7246   - Simproc "record_upd_simproc" for simplification of multiple updates added
  7247     (not enabled by default).
  7248   - Simproc "record_ex_sel_eq_simproc" to simplify EX x. sel r = x resp.
  7249     EX x. x = sel r to True (not enabled by default).
  7250   - Tactic "record_split_simp_tac" to split and simplify records added.
  7251 
  7252 * 'specification' command added, allowing for definition by
  7253   specification.  There is also an 'ax_specification' command that
  7254   introduces the new constants axiomatically.
  7255 
  7256 * arith(_tac) is now able to generate counterexamples for reals as well.
  7257 
  7258 * HOL-Algebra: new locale "ring" for non-commutative rings.
  7259 
  7260 * HOL-ex: InductiveInvariant_examples illustrates advanced recursive function
  7261   definitions, thanks to Sava Krsti\'{c} and John Matthews.
  7262 
  7263 * HOL-Matrix: a first theory for matrices in HOL with an application of
  7264   matrix theory to linear programming.
  7265 
  7266 * Unions and Intersections:
  7267   The latex output syntax of UN and INT has been changed
  7268   from "\Union x \in A. B" to "\Union_{x \in A} B"
  7269   i.e. the index formulae has become a subscript.
  7270   Similarly for "\Union x. B", and for \Inter instead of \Union.
  7271 
  7272 * Unions and Intersections over Intervals:
  7273   There is new short syntax "UN i<=n. A" for "UN i:{0..n}. A". There is
  7274   also an x-symbol version with subscripts "\<Union>\<^bsub>i <= n\<^esub>. A"
  7275   like in normal math, and corresponding versions for < and for intersection.
  7276 
  7277 * HOL/List: Ordering "lexico" is renamed "lenlex" and the standard
  7278   lexicographic dictonary ordering has been added as "lexord".
  7279 
  7280 * ML: the legacy theory structures Int and List have been removed. They had
  7281   conflicted with ML Basis Library structures having the same names.
  7282 
  7283 * 'refute' command added to search for (finite) countermodels.  Only works
  7284   for a fragment of HOL.  The installation of an external SAT solver is
  7285   highly recommended.  See "HOL/Refute.thy" for details.
  7286 
  7287 * 'quickcheck' command: Allows to find counterexamples by evaluating
  7288   formulae under an assignment of free variables to random values.
  7289   In contrast to 'refute', it can deal with inductive datatypes,
  7290   but cannot handle quantifiers. See "HOL/ex/Quickcheck_Examples.thy"
  7291   for examples.
  7292 
  7293 
  7294 *** HOLCF ***
  7295 
  7296 * Streams now come with concatenation and are part of the HOLCF image
  7297 
  7298 
  7299 
  7300 New in Isabelle2003 (May 2003)
  7301 ------------------------------
  7302 
  7303 *** General ***
  7304 
  7305 * Provers/simplifier:
  7306 
  7307   - Completely reimplemented method simp (ML: Asm_full_simp_tac):
  7308     Assumptions are now subject to complete mutual simplification,
  7309     not just from left to right. The simplifier now preserves
  7310     the order of assumptions.
  7311 
  7312     Potential INCOMPATIBILITY:
  7313 
  7314     -- simp sometimes diverges where the old version did
  7315        not, e.g. invoking simp on the goal
  7316 
  7317         [| P (f x); y = x; f x = f y |] ==> Q
  7318 
  7319        now gives rise to the infinite reduction sequence
  7320 
  7321         P(f x) --(f x = f y)--> P(f y) --(y = x)--> P(f x) --(f x = f y)--> ...
  7322 
  7323        Using "simp (asm_lr)" (ML: Asm_lr_simp_tac) instead often solves this
  7324        kind of problem.
  7325 
  7326     -- Tactics combining classical reasoner and simplification (such as auto)
  7327        are also affected by this change, because many of them rely on
  7328        simp. They may sometimes diverge as well or yield a different numbers
  7329        of subgoals. Try to use e.g. force, fastsimp, or safe instead of auto
  7330        in case of problems. Sometimes subsequent calls to the classical
  7331        reasoner will fail because a preceeding call to the simplifier too
  7332        eagerly simplified the goal, e.g. deleted redundant premises.
  7333 
  7334   - The simplifier trace now shows the names of the applied rewrite rules
  7335 
  7336   - You can limit the number of recursive invocations of the simplifier
  7337     during conditional rewriting (where the simplifie tries to solve the
  7338     conditions before applying the rewrite rule):
  7339     ML "simp_depth_limit := n"
  7340     where n is an integer. Thus you can force termination where previously
  7341     the simplifier would diverge.
  7342 
  7343   - Accepts free variables as head terms in congruence rules.  Useful in Isar.
  7344 
  7345   - No longer aborts on failed congruence proof.  Instead, the
  7346     congruence is ignored.
  7347 
  7348 * Pure: New generic framework for extracting programs from constructive
  7349   proofs. See HOL/Extraction.thy for an example instantiation, as well
  7350   as HOL/Extraction for some case studies.
  7351 
  7352 * Pure: The main goal of the proof state is no longer shown by default, only
  7353 the subgoals. This behaviour is controlled by a new flag.
  7354    PG menu: Isabelle/Isar -> Settings -> Show Main Goal
  7355 (ML: Proof.show_main_goal).
  7356 
  7357 * Pure: You can find all matching introduction rules for subgoal 1, i.e. all
  7358 rules whose conclusion matches subgoal 1:
  7359       PG menu: Isabelle/Isar -> Show me -> matching rules
  7360 The rules are ordered by how closely they match the subgoal.
  7361 In particular, rules that solve a subgoal outright are displayed first
  7362 (or rather last, the way they are printed).
  7363 (ML: ProofGeneral.print_intros())
  7364 
  7365 * Pure: New flag trace_unify_fail causes unification to print
  7366 diagnostic information (PG: in trace buffer) when it fails. This is
  7367 useful for figuring out why single step proofs like rule, erule or
  7368 assumption failed.
  7369 
  7370 * Pure: Locale specifications now produce predicate definitions
  7371 according to the body of text (covering assumptions modulo local
  7372 definitions); predicate "loc_axioms" covers newly introduced text,
  7373 while "loc" is cumulative wrt. all included locale expressions; the
  7374 latter view is presented only on export into the global theory
  7375 context; potential INCOMPATIBILITY, use "(open)" option to fall back
  7376 on the old view without predicates;
  7377 
  7378 * Pure: predefined locales "var" and "struct" are useful for sharing
  7379 parameters (as in CASL, for example); just specify something like
  7380 ``var x + var y + struct M'' as import;
  7381 
  7382 * Pure: improved thms_containing: proper indexing of facts instead of
  7383 raw theorems; check validity of results wrt. current name space;
  7384 include local facts of proof configuration (also covers active
  7385 locales), cover fixed variables in index; may use "_" in term
  7386 specification; an optional limit for the number of printed facts may
  7387 be given (the default is 40);
  7388 
  7389 * Pure: disallow duplicate fact bindings within new-style theory files
  7390 (batch-mode only);
  7391 
  7392 * Provers: improved induct method: assumptions introduced by case
  7393 "foo" are split into "foo.hyps" (from the rule) and "foo.prems" (from
  7394 the goal statement); "foo" still refers to all facts collectively;
  7395 
  7396 * Provers: the function blast.overloaded has been removed: all constants
  7397 are regarded as potentially overloaded, which improves robustness in exchange
  7398 for slight decrease in efficiency;
  7399 
  7400 * Provers/linorder: New generic prover for transitivity reasoning over
  7401 linear orders.  Note: this prover is not efficient!
  7402 
  7403 * Isar: preview of problems to finish 'show' now produce an error
  7404 rather than just a warning (in interactive mode);
  7405 
  7406 
  7407 *** HOL ***
  7408 
  7409 * arith(_tac)
  7410 
  7411  - Produces a counter example if it cannot prove a goal.
  7412    Note that the counter example may be spurious if the goal is not a formula
  7413    of quantifier-free linear arithmetic.
  7414    In ProofGeneral the counter example appears in the trace buffer.
  7415 
  7416  - Knows about div k and mod k where k is a numeral of type nat or int.
  7417 
  7418  - Calls full Presburger arithmetic (by Amine Chaieb) if quantifier-free
  7419    linear arithmetic fails. This takes account of quantifiers and divisibility.
  7420    Presburger arithmetic can also be called explicitly via presburger(_tac).
  7421 
  7422 * simp's arithmetic capabilities have been enhanced a bit: it now
  7423 takes ~= in premises into account (by performing a case split);
  7424 
  7425 * simp reduces "m*(n div m) + n mod m" to n, even if the two summands
  7426 are distributed over a sum of terms;
  7427 
  7428 * New tactic "trans_tac" and method "trans" instantiate
  7429 Provers/linorder.ML for axclasses "order" and "linorder" (predicates
  7430 "<=", "<" and "=").
  7431 
  7432 * function INCOMPATIBILITIES: Pi-sets have been redefined and moved from main
  7433 HOL to Library/FuncSet; constant "Fun.op o" is now called "Fun.comp";
  7434 
  7435 * 'typedef' command has new option "open" to suppress the set
  7436 definition;
  7437 
  7438 * functions Min and Max on finite sets have been introduced (theory
  7439 Finite_Set);
  7440 
  7441 * attribute [symmetric] now works for relations as well; it turns
  7442 (x,y) : R^-1 into (y,x) : R, and vice versa;
  7443 
  7444 * induct over a !!-quantified statement (say !!x1..xn):
  7445   each "case" automatically performs "fix x1 .. xn" with exactly those names.
  7446 
  7447 * Map: `empty' is no longer a constant but a syntactic abbreviation for
  7448 %x. None. Warning: empty_def now refers to the previously hidden definition
  7449 of the empty set.
  7450 
  7451 * Algebra: formalization of classical algebra.  Intended as base for
  7452 any algebraic development in Isabelle.  Currently covers group theory
  7453 (up to Sylow's theorem) and ring theory (Universal Property of
  7454 Univariate Polynomials).  Contributions welcome;
  7455 
  7456 * GroupTheory: deleted, since its material has been moved to Algebra;
  7457 
  7458 * Complex: new directory of the complex numbers with numeric constants,
  7459 nonstandard complex numbers, and some complex analysis, standard and
  7460 nonstandard (Jacques Fleuriot);
  7461 
  7462 * HOL-Complex: new image for analysis, replacing HOL-Real and HOL-Hyperreal;
  7463 
  7464 * Hyperreal: introduced Gauge integration and hyperreal logarithms (Jacques
  7465 Fleuriot);
  7466 
  7467 * Real/HahnBanach: updated and adapted to locales;
  7468 
  7469 * NumberTheory: added Gauss's law of quadratic reciprocity (by Avigad,
  7470 Gray and Kramer);
  7471 
  7472 * UNITY: added the Meier-Sanders theory of progress sets;
  7473 
  7474 * MicroJava: bytecode verifier and lightweight bytecode verifier
  7475 as abstract algorithms, instantiated to the JVM;
  7476 
  7477 * Bali: Java source language formalization. Type system, operational
  7478 semantics, axiomatic semantics. Supported language features:
  7479 classes, interfaces, objects,virtual methods, static methods,
  7480 static/instance fields, arrays, access modifiers, definite
  7481 assignment, exceptions.
  7482 
  7483 
  7484 *** ZF ***
  7485 
  7486 * ZF/Constructible: consistency proof for AC (Gdel's constructible
  7487 universe, etc.);
  7488 
  7489 * Main ZF: virtually all theories converted to new-style format;
  7490 
  7491 
  7492 *** ML ***
  7493 
  7494 * Pure: Tactic.prove provides sane interface for internal proofs;
  7495 omits the infamous "standard" operation, so this is more appropriate
  7496 than prove_goalw_cterm in many situations (e.g. in simprocs);
  7497 
  7498 * Pure: improved error reporting of simprocs;
  7499 
  7500 * Provers: Simplifier.simproc(_i) provides sane interface for setting
  7501 up simprocs;
  7502 
  7503 
  7504 *** Document preparation ***
  7505 
  7506 * uses \par instead of \\ for line breaks in theory text. This may
  7507 shift some page breaks in large documents. To get the old behaviour
  7508 use \renewcommand{\isanewline}{\mbox{}\\\mbox{}} in root.tex.
  7509 
  7510 * minimized dependencies of isabelle.sty and isabellesym.sty on
  7511 other packages
  7512 
  7513 * \<euro> now needs package babel/greek instead of marvosym (which
  7514 broke \Rightarrow)
  7515 
  7516 * normal size for \<zero>...\<nine> (uses \mathbf instead of
  7517 textcomp package)
  7518 
  7519 
  7520 
  7521 New in Isabelle2002 (March 2002)
  7522 --------------------------------
  7523 
  7524 *** Document preparation ***
  7525 
  7526 * greatly simplified document preparation setup, including more
  7527 graceful interpretation of isatool usedir -i/-d/-D options, and more
  7528 instructive isatool mkdir; users should basically be able to get
  7529 started with "isatool mkdir HOL Test && isatool make"; alternatively,
  7530 users may run a separate document processing stage manually like this:
  7531 "isatool usedir -D output HOL Test && isatool document Test/output";
  7532 
  7533 * theory dependency graph may now be incorporated into documents;
  7534 isatool usedir -g true will produce session_graph.eps/.pdf for use
  7535 with \includegraphics of LaTeX;
  7536 
  7537 * proper spacing of consecutive markup elements, especially text
  7538 blocks after section headings;
  7539 
  7540 * support bold style (for single symbols only), input syntax is like
  7541 this: "\<^bold>\<alpha>" or "\<^bold>A";
  7542 
  7543 * \<bullet> is now output as bold \cdot by default, which looks much
  7544 better in printed text;
  7545 
  7546 * added default LaTeX bindings for \<tturnstile> and \<TTurnstile>;
  7547 note that these symbols are currently unavailable in Proof General /
  7548 X-Symbol; new symbols \<zero>, \<one>, ..., \<nine>, and \<euro>;
  7549 
  7550 * isatool latex no longer depends on changed TEXINPUTS, instead
  7551 isatool document copies the Isabelle style files to the target
  7552 location;
  7553 
  7554 
  7555 *** Isar ***
  7556 
  7557 * Pure/Provers: improved proof by cases and induction;
  7558   - 'case' command admits impromptu naming of parameters (such as
  7559     "case (Suc n)");
  7560   - 'induct' method divinates rule instantiation from the inductive
  7561     claim; no longer requires excessive ?P bindings for proper
  7562     instantiation of cases;
  7563   - 'induct' method properly enumerates all possibilities of set/type
  7564     rules; as a consequence facts may be also passed through *type*
  7565     rules without further ado;
  7566   - 'induct' method now derives symbolic cases from the *rulified*
  7567     rule (before it used to rulify cases stemming from the internal
  7568     atomized version); this means that the context of a non-atomic
  7569     statement becomes is included in the hypothesis, avoiding the
  7570     slightly cumbersome show "PROP ?case" form;
  7571   - 'induct' may now use elim-style induction rules without chaining
  7572     facts, using ``missing'' premises from the goal state; this allows
  7573     rules stemming from inductive sets to be applied in unstructured
  7574     scripts, while still benefitting from proper handling of non-atomic
  7575     statements; NB: major inductive premises need to be put first, all
  7576     the rest of the goal is passed through the induction;
  7577   - 'induct' proper support for mutual induction involving non-atomic
  7578     rule statements (uses the new concept of simultaneous goals, see
  7579     below);
  7580   - append all possible rule selections, but only use the first
  7581     success (no backtracking);
  7582   - removed obsolete "(simplified)" and "(stripped)" options of methods;
  7583   - undeclared rule case names default to numbers 1, 2, 3, ...;
  7584   - added 'print_induct_rules' (covered by help item in recent Proof
  7585     General versions);
  7586   - moved induct/cases attributes to Pure, methods to Provers;
  7587   - generic method setup instantiated for FOL and HOL;
  7588 
  7589 * Pure: support multiple simultaneous goal statements, for example
  7590 "have a: A and b: B" (same for 'theorem' etc.); being a pure
  7591 meta-level mechanism, this acts as if several individual goals had
  7592 been stated separately; in particular common proof methods need to be
  7593 repeated in order to cover all claims; note that a single elimination
  7594 step is *not* sufficient to establish the two conjunctions, so this
  7595 fails:
  7596 
  7597   assume "A & B" then have A and B ..   (*".." fails*)
  7598 
  7599 better use "obtain" in situations as above; alternative refer to
  7600 multi-step methods like 'auto', 'simp_all', 'blast+' etc.;
  7601 
  7602 * Pure: proper integration with ``locales''; unlike the original
  7603 version by Florian Kammller, Isar locales package high-level proof
  7604 contexts rather than raw logical ones (e.g. we admit to include
  7605 attributes everywhere); operations on locales include merge and
  7606 rename; support for implicit arguments (``structures''); simultaneous
  7607 type-inference over imports and text; see also HOL/ex/Locales.thy for
  7608 some examples;
  7609 
  7610 * Pure: the following commands have been ``localized'', supporting a
  7611 target locale specification "(in name)": 'lemma', 'theorem',
  7612 'corollary', 'lemmas', 'theorems', 'declare'; the results will be
  7613 stored both within the locale and at the theory level (exported and
  7614 qualified by the locale name);
  7615 
  7616 * Pure: theory goals may now be specified in ``long'' form, with
  7617 ad-hoc contexts consisting of arbitrary locale elements. for example
  7618 ``lemma foo: fixes x assumes "A x" shows "B x"'' (local syntax and
  7619 definitions may be given, too); the result is a meta-level rule with
  7620 the context elements being discharged in the obvious way;
  7621 
  7622 * Pure: new proof command 'using' allows to augment currently used
  7623 facts after a goal statement ('using' is syntactically analogous to
  7624 'apply', but acts on the goal's facts only); this allows chained facts
  7625 to be separated into parts given before and after a claim, as in
  7626 ``from a and b have C using d and e <proof>'';
  7627 
  7628 * Pure: renamed "antecedent" case to "rule_context";
  7629 
  7630 * Pure: new 'judgment' command records explicit information about the
  7631 object-logic embedding (used by several tools internally); no longer
  7632 use hard-wired "Trueprop";
  7633 
  7634 * Pure: added 'corollary' command;
  7635 
  7636 * Pure: fixed 'token_translation' command;
  7637 
  7638 * Pure: removed obsolete 'exported' attribute;
  7639 
  7640 * Pure: dummy pattern "_" in is/let is now automatically lifted over
  7641 bound variables: "ALL x. P x --> Q x" (is "ALL x. _ --> ?C x")
  7642 supersedes more cumbersome ... (is "ALL x. _ x --> ?C x");
  7643 
  7644 * Pure: method 'atomize' presents local goal premises as object-level
  7645 statements (atomic meta-level propositions); setup controlled via
  7646 rewrite rules declarations of 'atomize' attribute; example
  7647 application: 'induct' method with proper rule statements in improper
  7648 proof *scripts*;
  7649 
  7650 * Pure: emulation of instantiation tactics (rule_tac, cut_tac, etc.)
  7651 now consider the syntactic context of assumptions, giving a better
  7652 chance to get type-inference of the arguments right (this is
  7653 especially important for locales);
  7654 
  7655 * Pure: "sorry" no longer requires quick_and_dirty in interactive
  7656 mode;
  7657 
  7658 * Pure/obtain: the formal conclusion "thesis", being marked as
  7659 ``internal'', may no longer be reference directly in the text;
  7660 potential INCOMPATIBILITY, may need to use "?thesis" in rare
  7661 situations;
  7662 
  7663 * Pure: generic 'sym' attribute which declares a rule both as pure
  7664 'elim?' and for the 'symmetric' operation;
  7665 
  7666 * Pure: marginal comments ``--'' may now occur just anywhere in the
  7667 text; the fixed correlation with particular command syntax has been
  7668 discontinued;
  7669 
  7670 * Pure: new method 'rules' is particularly well-suited for proof
  7671 search in intuitionistic logic; a bit slower than 'blast' or 'fast',
  7672 but often produces more compact proof terms with less detours;
  7673 
  7674 * Pure/Provers/classical: simplified integration with pure rule
  7675 attributes and methods; the classical "intro?/elim?/dest?"
  7676 declarations coincide with the pure ones; the "rule" method no longer
  7677 includes classically swapped intros; "intro" and "elim" methods no
  7678 longer pick rules from the context; also got rid of ML declarations
  7679 AddXIs/AddXEs/AddXDs; all of this has some potential for
  7680 INCOMPATIBILITY;
  7681 
  7682 * Provers/classical: attribute 'swapped' produces classical inversions
  7683 of introduction rules;
  7684 
  7685 * Provers/simplifier: 'simplified' attribute may refer to explicit
  7686 rules instead of full simplifier context; 'iff' attribute handles
  7687 conditional rules;
  7688 
  7689 * HOL: 'typedef' now allows alternative names for Rep/Abs morphisms;
  7690 
  7691 * HOL: 'recdef' now fails on unfinished automated proofs, use
  7692 "(permissive)" option to recover old behavior;
  7693 
  7694 * HOL: 'inductive' no longer features separate (collective) attributes
  7695 for 'intros' (was found too confusing);
  7696 
  7697 * HOL: properly declared induction rules less_induct and
  7698 wf_induct_rule;
  7699 
  7700 
  7701 *** HOL ***
  7702 
  7703 * HOL: moved over to sane numeral syntax; the new policy is as
  7704 follows:
  7705 
  7706   - 0 and 1 are polymorphic constants, which are defined on any
  7707   numeric type (nat, int, real etc.);
  7708 
  7709   - 2, 3, 4, ... and -1, -2, -3, ... are polymorphic numerals, based
  7710   binary representation internally;
  7711 
  7712   - type nat has special constructor Suc, and generally prefers Suc 0
  7713   over 1::nat and Suc (Suc 0) over 2::nat;
  7714 
  7715 This change may cause significant problems of INCOMPATIBILITY; here
  7716 are some hints on converting existing sources:
  7717 
  7718   - due to the new "num" token, "-0" and "-1" etc. are now atomic
  7719   entities, so expressions involving "-" (unary or binary minus) need
  7720   to be spaced properly;
  7721 
  7722   - existing occurrences of "1" may need to be constraint "1::nat" or
  7723   even replaced by Suc 0; similar for old "2";
  7724 
  7725   - replace "#nnn" by "nnn", and "#-nnn" by "-nnn";
  7726 
  7727   - remove all special provisions on numerals in proofs;
  7728 
  7729 * HOL: simp rules nat_number expand numerals on nat to Suc/0
  7730 representation (depends on bin_arith_simps in the default context);
  7731 
  7732 * HOL: symbolic syntax for x^2 (numeral 2);
  7733 
  7734 * HOL: the class of all HOL types is now called "type" rather than
  7735 "term"; INCOMPATIBILITY, need to adapt references to this type class
  7736 in axclass/classes, instance/arities, and (usually rare) occurrences
  7737 in typings (of consts etc.); internally the class is called
  7738 "HOL.type", ML programs should refer to HOLogic.typeS;
  7739 
  7740 * HOL/record package improvements:
  7741   - new derived operations "fields" to build a partial record section,
  7742     "extend" to promote a fixed record to a record scheme, and
  7743     "truncate" for the reverse; cf. theorems "xxx.defs", which are *not*
  7744     declared as simp by default;
  7745   - shared operations ("more", "fields", etc.) now need to be always
  7746     qualified) --- potential INCOMPATIBILITY;
  7747   - removed "make_scheme" operations (use "make" with "extend") --
  7748     INCOMPATIBILITY;
  7749   - removed "more" class (simply use "term") -- INCOMPATIBILITY;
  7750   - provides cases/induct rules for use with corresponding Isar
  7751     methods (for concrete records, record schemes, concrete more
  7752     parts, and schematic more parts -- in that order);
  7753   - internal definitions directly based on a light-weight abstract
  7754     theory of product types over typedef rather than datatype;
  7755 
  7756 * HOL: generic code generator for generating executable ML code from
  7757 specifications; specific support for HOL constructs such as inductive
  7758 datatypes and sets, as well as recursive functions; can be invoked
  7759 via 'generate_code' theory section;
  7760 
  7761 * HOL: canonical cases/induct rules for n-tuples (n = 3..7);
  7762 
  7763 * HOL: consolidated and renamed several theories.  In particular:
  7764         Ord.thy has been absorbed into HOL.thy
  7765         String.thy has been absorbed into List.thy
  7766 
  7767 * HOL: concrete setsum syntax "\<Sum>i:A. b" == "setsum (%i. b) A"
  7768 (beware of argument permutation!);
  7769 
  7770 * HOL: linorder_less_split superseded by linorder_cases;
  7771 
  7772 * HOL/List: "nodups" renamed to "distinct";
  7773 
  7774 * HOL: added "The" definite description operator; move Hilbert's "Eps"
  7775 to peripheral theory "Hilbert_Choice"; some INCOMPATIBILITIES:
  7776   - Ex_def has changed, now need to use some_eq_ex
  7777 
  7778 * HOL: made split_all_tac safe; EXISTING PROOFS MAY FAIL OR LOOP, so
  7779 in this (rare) case use:
  7780 
  7781   delSWrapper "split_all_tac"
  7782   addSbefore ("unsafe_split_all_tac", unsafe_split_all_tac)
  7783 
  7784 * HOL: added safe wrapper "split_conv_tac" to claset; EXISTING PROOFS
  7785 MAY FAIL;
  7786 
  7787 * HOL: introduced f^n = f o ... o f; warning: due to the limits of
  7788 Isabelle's type classes, ^ on functions and relations has too general
  7789 a domain, namely ('a * 'b) set and 'a => 'b; this means that it may be
  7790 necessary to attach explicit type constraints;
  7791 
  7792 * HOL/Relation: the prefix name of the infix "O" has been changed from
  7793 "comp" to "rel_comp"; INCOMPATIBILITY: a few theorems have been
  7794 renamed accordingly (eg "compI" -> "rel_compI").
  7795 
  7796 * HOL: syntax translations now work properly with numerals and records
  7797 expressions;
  7798 
  7799 * HOL: bounded abstraction now uses syntax "%" / "\<lambda>" instead
  7800 of "lam" -- INCOMPATIBILITY;
  7801 
  7802 * HOL: got rid of some global declarations (potential INCOMPATIBILITY
  7803 for ML tools): const "()" renamed "Product_Type.Unity", type "unit"
  7804 renamed "Product_Type.unit";
  7805 
  7806 * HOL: renamed rtrancl_into_rtrancl2 to converse_rtrancl_into_rtrancl
  7807 
  7808 * HOL: removed obsolete theorem "optionE" (use "option.exhaust", or
  7809 the "cases" method);
  7810 
  7811 * HOL/GroupTheory: group theory examples including Sylow's theorem (by
  7812 Florian Kammller);
  7813 
  7814 * HOL/IMP: updated and converted to new-style theory format; several
  7815 parts turned into readable document, with proper Isar proof texts and
  7816 some explanations (by Gerwin Klein);
  7817 
  7818 * HOL-Real: added Complex_Numbers (by Gertrud Bauer);
  7819 
  7820 * HOL-Hyperreal is now a logic image;
  7821 
  7822 
  7823 *** HOLCF ***
  7824 
  7825 * Isar: consts/constdefs supports mixfix syntax for continuous
  7826 operations;
  7827 
  7828 * Isar: domain package adapted to new-style theory format, e.g. see
  7829 HOLCF/ex/Dnat.thy;
  7830 
  7831 * theory Lift: proper use of rep_datatype lift instead of ML hacks --
  7832 potential INCOMPATIBILITY; now use plain induct_tac instead of former
  7833 lift.induct_tac, always use UU instead of Undef;
  7834 
  7835 * HOLCF/IMP: updated and converted to new-style theory;
  7836 
  7837 
  7838 *** ZF ***
  7839 
  7840 * Isar: proper integration of logic-specific tools and packages,
  7841 including theory commands '(co)inductive', '(co)datatype',
  7842 'rep_datatype', 'inductive_cases', as well as methods 'ind_cases',
  7843 'induct_tac', 'case_tac', and 'typecheck' (with attribute 'TC');
  7844 
  7845 * theory Main no longer includes AC; for the Axiom of Choice, base
  7846 your theory on Main_ZFC;
  7847 
  7848 * the integer library now covers quotients and remainders, with many
  7849 laws relating division to addition, multiplication, etc.;
  7850 
  7851 * ZF/UNITY: Chandy and Misra's UNITY is now available in ZF, giving a
  7852 typeless version of the formalism;
  7853 
  7854 * ZF/AC, Coind, IMP, Resid: updated and converted to new-style theory
  7855 format;
  7856 
  7857 * ZF/Induct: new directory for examples of inductive definitions,
  7858 including theory Multiset for multiset orderings; converted to
  7859 new-style theory format;
  7860 
  7861 * ZF: many new theorems about lists, ordinals, etc.;
  7862 
  7863 
  7864 *** General ***
  7865 
  7866 * Pure/kernel: meta-level proof terms (by Stefan Berghofer); reference
  7867 variable proof controls level of detail: 0 = no proofs (only oracle
  7868 dependencies), 1 = lemma dependencies, 2 = compact proof terms; see
  7869 also ref manual for further ML interfaces;
  7870 
  7871 * Pure/axclass: removed obsolete ML interface
  7872 goal_subclass/goal_arity;
  7873 
  7874 * Pure/syntax: new token syntax "num" for plain numerals (without "#"
  7875 of "xnum"); potential INCOMPATIBILITY, since -0, -1 etc. are now
  7876 separate tokens, so expressions involving minus need to be spaced
  7877 properly;
  7878 
  7879 * Pure/syntax: support non-oriented infixes, using keyword "infix"
  7880 rather than "infixl" or "infixr";
  7881 
  7882 * Pure/syntax: concrete syntax for dummy type variables admits genuine
  7883 sort constraint specifications in type inference; e.g. "x::_::foo"
  7884 ensures that the type of "x" is of sort "foo" (but not necessarily a
  7885 type variable);
  7886 
  7887 * Pure/syntax: print modes "type_brackets" and "no_type_brackets"
  7888 control output of nested => (types); the default behavior is
  7889 "type_brackets";
  7890 
  7891 * Pure/syntax: builtin parse translation for "_constify" turns valued
  7892 tokens into AST constants;
  7893 
  7894 * Pure/syntax: prefer later declarations of translations and print
  7895 translation functions; potential INCOMPATIBILITY: need to reverse
  7896 multiple declarations for same syntax element constant;
  7897 
  7898 * Pure/show_hyps reset by default (in accordance to existing Isar
  7899 practice);
  7900 
  7901 * Provers/classical: renamed addaltern to addafter, addSaltern to
  7902 addSafter;
  7903 
  7904 * Provers/clasimp: ``iff'' declarations now handle conditional rules
  7905 as well;
  7906 
  7907 * system: tested support for MacOS X; should be able to get Isabelle +
  7908 Proof General to work in a plain Terminal after installing Poly/ML
  7909 (e.g. from the Isabelle distribution area) and GNU bash alone
  7910 (e.g. from http://www.apple.com); full X11, XEmacs and X-Symbol
  7911 support requires further installations, e.g. from
  7912 http://fink.sourceforge.net/);
  7913 
  7914 * system: support Poly/ML 4.1.1 (able to manage larger heaps);
  7915 
  7916 * system: reduced base memory usage by Poly/ML (approx. 20 MB instead
  7917 of 40 MB), cf. ML_OPTIONS;
  7918 
  7919 * system: Proof General keywords specification is now part of the
  7920 Isabelle distribution (see etc/isar-keywords.el);
  7921 
  7922 * system: support for persistent Proof General sessions (refrain from
  7923 outdating all loaded theories on startup); user may create writable
  7924 logic images like this: ``isabelle -q HOL Test'';
  7925 
  7926 * system: smart selection of Isabelle process versus Isabelle
  7927 interface, accommodates case-insensitive file systems (e.g. HFS+); may
  7928 run both "isabelle" and "Isabelle" even if file names are badly
  7929 damaged (executable inspects the case of the first letter of its own
  7930 name); added separate "isabelle-process" and "isabelle-interface";
  7931 
  7932 * system: refrain from any attempt at filtering input streams; no
  7933 longer support ``8bit'' encoding of old isabelle font, instead proper
  7934 iso-latin characters may now be used; the related isatools
  7935 "symbolinput" and "nonascii" have disappeared as well;
  7936 
  7937 * system: removed old "xterm" interface (the print modes "xterm" and
  7938 "xterm_color" are still available for direct use in a suitable
  7939 terminal);
  7940 
  7941 
  7942 
  7943 New in Isabelle99-2 (February 2001)
  7944 -----------------------------------
  7945 
  7946 *** Overview of INCOMPATIBILITIES ***
  7947 
  7948 * HOL: please note that theories in the Library and elsewhere often use the
  7949 new-style (Isar) format; to refer to their theorems in an ML script you must
  7950 bind them to ML identifers by e.g.      val thm_name = thm "thm_name";
  7951 
  7952 * HOL: inductive package no longer splits induction rule aggressively,
  7953 but only as far as specified by the introductions given; the old
  7954 format may be recovered via ML function complete_split_rule or attribute
  7955 'split_rule (complete)';
  7956 
  7957 * HOL: induct renamed to lfp_induct, lfp_Tarski to lfp_unfold,
  7958 gfp_Tarski to gfp_unfold;
  7959 
  7960 * HOL: contrapos, contrapos2 renamed to contrapos_nn, contrapos_pp;
  7961 
  7962 * HOL: infix "dvd" now has priority 50 rather than 70 (because it is a
  7963 relation); infix "^^" has been renamed "``"; infix "``" has been
  7964 renamed "`"; "univalent" has been renamed "single_valued";
  7965 
  7966 * HOL/Real: "rinv" and "hrinv" replaced by overloaded "inverse"
  7967 operation;
  7968 
  7969 * HOLCF: infix "`" has been renamed "$"; the symbol syntax is \<cdot>;
  7970 
  7971 * Isar: 'obtain' no longer declares "that" fact as simp/intro;
  7972 
  7973 * Isar/HOL: method 'induct' now handles non-atomic goals; as a
  7974 consequence, it is no longer monotonic wrt. the local goal context
  7975 (which is now passed through the inductive cases);
  7976 
  7977 * Document preparation: renamed standard symbols \<ll> to \<lless> and
  7978 \<gg> to \<ggreater>;
  7979 
  7980 
  7981 *** Document preparation ***
  7982 
  7983 * \isabellestyle{NAME} selects version of Isabelle output (currently
  7984 available: are "it" for near math-mode best-style output, "sl" for
  7985 slanted text style, and "tt" for plain type-writer; if no
  7986 \isabellestyle command is given, output is according to slanted
  7987 type-writer);
  7988 
  7989 * support sub/super scripts (for single symbols only), input syntax is
  7990 like this: "A\<^sup>*" or "A\<^sup>\<star>";
  7991 
  7992 * some more standard symbols; see Appendix A of the system manual for
  7993 the complete list of symbols defined in isabellesym.sty;
  7994 
  7995 * improved isabelle style files; more abstract symbol implementation
  7996 (should now use \isamath{...} and \isatext{...} in custom symbol
  7997 definitions);
  7998 
  7999 * antiquotation @{goals} and @{subgoals} for output of *dynamic* goals
  8000 state; Note that presentation of goal states does not conform to
  8001 actual human-readable proof documents.  Please do not include goal
  8002 states into document output unless you really know what you are doing!
  8003 
  8004 * proper indentation of antiquoted output with proportional LaTeX
  8005 fonts;
  8006 
  8007 * no_document ML operator temporarily disables LaTeX document
  8008 generation;
  8009 
  8010 * isatool unsymbolize tunes sources for plain ASCII communication;
  8011 
  8012 
  8013 *** Isar ***
  8014 
  8015 * Pure: Isar now suffers initial goal statements to contain unbound
  8016 schematic variables (this does not conform to actual readable proof
  8017 documents, due to unpredictable outcome and non-compositional proof
  8018 checking); users who know what they are doing may use schematic goals
  8019 for Prolog-style synthesis of proven results;
  8020 
  8021 * Pure: assumption method (an implicit finishing) now handles actual
  8022 rules as well;
  8023 
  8024 * Pure: improved 'obtain' --- moved to Pure, insert "that" into
  8025 initial goal, declare "that" only as Pure intro (only for single
  8026 steps); the "that" rule assumption may now be involved in implicit
  8027 finishing, thus ".." becomes a feasible for trivial obtains;
  8028 
  8029 * Pure: default proof step now includes 'intro_classes'; thus trivial
  8030 instance proofs may be performed by "..";
  8031 
  8032 * Pure: ?thesis / ?this / "..." now work for pure meta-level
  8033 statements as well;
  8034 
  8035 * Pure: more robust selection of calculational rules;
  8036 
  8037 * Pure: the builtin notion of 'finished' goal now includes the ==-refl
  8038 rule (as well as the assumption rule);
  8039 
  8040 * Pure: 'thm_deps' command visualizes dependencies of theorems and
  8041 lemmas, using the graph browser tool;
  8042 
  8043 * Pure: predict failure of "show" in interactive mode;
  8044 
  8045 * Pure: 'thms_containing' now takes actual terms as arguments;
  8046 
  8047 * HOL: improved method 'induct' --- now handles non-atomic goals
  8048 (potential INCOMPATIBILITY); tuned error handling;
  8049 
  8050 * HOL: cases and induct rules now provide explicit hints about the
  8051 number of facts to be consumed (0 for "type" and 1 for "set" rules);
  8052 any remaining facts are inserted into the goal verbatim;
  8053 
  8054 * HOL: local contexts (aka cases) may now contain term bindings as
  8055 well; the 'cases' and 'induct' methods new provide a ?case binding for
  8056 the result to be shown in each case;
  8057 
  8058 * HOL: added 'recdef_tc' command;
  8059 
  8060 * isatool convert assists in eliminating legacy ML scripts;
  8061 
  8062 
  8063 *** HOL ***
  8064 
  8065 * HOL/Library: a collection of generic theories to be used together
  8066 with main HOL; the theory loader path already includes this directory
  8067 by default; the following existing theories have been moved here:
  8068 HOL/Induct/Multiset, HOL/Induct/Acc (as Accessible_Part), HOL/While
  8069 (as While_Combinator), HOL/Lex/Prefix (as List_Prefix);
  8070 
  8071 * HOL/Unix: "Some aspects of Unix file-system security", a typical
  8072 modelling and verification task performed in Isabelle/HOL +
  8073 Isabelle/Isar + Isabelle document preparation (by Markus Wenzel).
  8074 
  8075 * HOL/Algebra: special summation operator SUM no longer exists, it has
  8076 been replaced by setsum; infix 'assoc' now has priority 50 (like
  8077 'dvd'); axiom 'one_not_zero' has been moved from axclass 'ring' to
  8078 'domain', this makes the theory consistent with mathematical
  8079 literature;
  8080 
  8081 * HOL basics: added overloaded operations "inverse" and "divide"
  8082 (infix "/"), syntax for generic "abs" operation, generic summation
  8083 operator \<Sum>;
  8084 
  8085 * HOL/typedef: simplified package, provide more useful rules (see also
  8086 HOL/subset.thy);
  8087 
  8088 * HOL/datatype: induction rule for arbitrarily branching datatypes is
  8089 now expressed as a proper nested rule (old-style tactic scripts may
  8090 require atomize_strip_tac to cope with non-atomic premises);
  8091 
  8092 * HOL: renamed theory "Prod" to "Product_Type", renamed "split" rule
  8093 to "split_conv" (old name still available for compatibility);
  8094 
  8095 * HOL: improved concrete syntax for strings (e.g. allows translation
  8096 rules with string literals);
  8097 
  8098 * HOL-Real-Hyperreal: this extends HOL-Real with the hyperreals
  8099  and Fleuriot's mechanization of analysis, including the transcendental
  8100  functions for the reals;
  8101 
  8102 * HOL/Real, HOL/Hyperreal: improved arithmetic simplification;
  8103 
  8104 
  8105 *** CTT ***
  8106 
  8107 * CTT: x-symbol support for Pi, Sigma, -->, : (membership); note that
  8108 "lam" is displayed as TWO lambda-symbols
  8109 
  8110 * CTT: theory Main now available, containing everything (that is, Bool
  8111 and Arith);
  8112 
  8113 
  8114 *** General ***
  8115 
  8116 * Pure: the Simplifier has been implemented properly as a derived rule
  8117 outside of the actual kernel (at last!); the overall performance
  8118 penalty in practical applications is about 50%, while reliability of
  8119 the Isabelle inference kernel has been greatly improved;
  8120 
  8121 * print modes "brackets" and "no_brackets" control output of nested =>
  8122 (types) and ==> (props); the default behaviour is "brackets";
  8123 
  8124 * Provers: fast_tac (and friends) now handle actual object-logic rules
  8125 as assumptions as well;
  8126 
  8127 * system: support Poly/ML 4.0;
  8128 
  8129 * system: isatool install handles KDE version 1 or 2;
  8130 
  8131 
  8132 
  8133 New in Isabelle99-1 (October 2000)
  8134 ----------------------------------
  8135 
  8136 *** Overview of INCOMPATIBILITIES ***
  8137 
  8138 * HOL: simplification of natural numbers is much changed; to partly
  8139 recover the old behaviour (e.g. to prevent n+n rewriting to #2*n)
  8140 issue the following ML commands:
  8141 
  8142   Delsimprocs Nat_Numeral_Simprocs.cancel_numerals;
  8143   Delsimprocs [Nat_Numeral_Simprocs.combine_numerals];
  8144 
  8145 * HOL: simplification no longer dives into case-expressions; this is
  8146 controlled by "t.weak_case_cong" for each datatype t;
  8147 
  8148 * HOL: nat_less_induct renamed to less_induct;
  8149 
  8150 * HOL: systematic renaming of the SOME (Eps) rules, may use isatool
  8151 fixsome to patch .thy and .ML sources automatically;
  8152 
  8153   select_equality  -> some_equality
  8154   select_eq_Ex     -> some_eq_ex
  8155   selectI2EX       -> someI2_ex
  8156   selectI2         -> someI2
  8157   selectI          -> someI
  8158   select1_equality -> some1_equality
  8159   Eps_sym_eq       -> some_sym_eq_trivial
  8160   Eps_eq           -> some_eq_trivial
  8161 
  8162 * HOL: exhaust_tac on datatypes superceded by new generic case_tac;
  8163 
  8164 * HOL: removed obsolete theorem binding expand_if (refer to split_if
  8165 instead);
  8166 
  8167 * HOL: the recursion equations generated by 'recdef' are now called
  8168 f.simps instead of f.rules;
  8169 
  8170 * HOL: qed_spec_mp now also handles bounded ALL as well;
  8171 
  8172 * HOL: 0 is now overloaded, so the type constraint ":: nat" may
  8173 sometimes be needed;
  8174 
  8175 * HOL: the constant for "f``x" is now "image" rather than "op ``";
  8176 
  8177 * HOL: the constant for "f-``x" is now "vimage" rather than "op -``";
  8178 
  8179 * HOL: the disjoint sum is now "<+>" instead of "Plus"; the cartesian
  8180 product is now "<*>" instead of "Times"; the lexicographic product is
  8181 now "<*lex*>" instead of "**";
  8182 
  8183 * HOL: theory Sexp is now in HOL/Induct examples (it used to be part
  8184 of main HOL, but was unused); better use HOL's datatype package;
  8185 
  8186 * HOL: removed "symbols" syntax for constant "override" of theory Map;
  8187 the old syntax may be recovered as follows:
  8188 
  8189   syntax (symbols)
  8190     override  :: "('a ~=> 'b) => ('a ~=> 'b) => ('a ~=> 'b)"
  8191       (infixl "\\<oplus>" 100)
  8192 
  8193 * HOL/Real: "rabs" replaced by overloaded "abs" function;
  8194 
  8195 * HOL/ML: even fewer consts are declared as global (see theories Ord,
  8196 Lfp, Gfp, WF); this only affects ML packages that refer to const names
  8197 internally;
  8198 
  8199 * HOL and ZF: syntax for quotienting wrt an equivalence relation
  8200 changed from A/r to A//r;
  8201 
  8202 * ZF: new treatment of arithmetic (nat & int) may break some old
  8203 proofs;
  8204 
  8205 * Isar: renamed some attributes (RS -> THEN, simplify -> simplified,
  8206 rulify -> rule_format, elimify -> elim_format, ...);
  8207 
  8208 * Isar/Provers: intro/elim/dest attributes changed; renamed
  8209 intro/intro!/intro!! flags to intro!/intro/intro? (in most cases, one
  8210 should have to change intro!! to intro? only); replaced "delrule" by
  8211 "rule del";
  8212 
  8213 * Isar/HOL: renamed "intrs" to "intros" in inductive definitions;
  8214 
  8215 * Provers: strengthened force_tac by using new first_best_tac;
  8216 
  8217 * LaTeX document preparation: several changes of isabelle.sty (see
  8218 lib/texinputs);
  8219 
  8220 
  8221 *** Document preparation ***
  8222 
  8223 * formal comments (text blocks etc.) in new-style theories may now
  8224 contain antiquotations of thm/prop/term/typ/text to be presented
  8225 according to latex print mode; concrete syntax is like this:
  8226 @{term[show_types] "f(x) = a + x"};
  8227 
  8228 * isatool mkdir provides easy setup of Isabelle session directories,
  8229 including proper document sources;
  8230 
  8231 * generated LaTeX sources are now deleted after successful run
  8232 (isatool document -c); may retain a copy somewhere else via -D option
  8233 of isatool usedir;
  8234 
  8235 * isatool usedir -D now lets isatool latex -o sty update the Isabelle
  8236 style files, achieving self-contained LaTeX sources and simplifying
  8237 LaTeX debugging;
  8238 
  8239 * old-style theories now produce (crude) LaTeX output as well;
  8240 
  8241 * browser info session directories are now self-contained (may be put
  8242 on WWW server seperately); improved graphs of nested sessions; removed
  8243 graph for 'all sessions';
  8244 
  8245 * several improvements in isabelle style files; \isabellestyle{it}
  8246 produces fake math mode output; \isamarkupheader is now \section by
  8247 default; see lib/texinputs/isabelle.sty etc.;
  8248 
  8249 
  8250 *** Isar ***
  8251 
  8252 * Isar/Pure: local results and corresponding term bindings are now
  8253 subject to Hindley-Milner polymorphism (similar to ML); this
  8254 accommodates incremental type-inference very nicely;
  8255 
  8256 * Isar/Pure: new derived language element 'obtain' supports
  8257 generalized existence reasoning;
  8258 
  8259 * Isar/Pure: new calculational elements 'moreover' and 'ultimately'
  8260 support accumulation of results, without applying any rules yet;
  8261 useful to collect intermediate results without explicit name
  8262 references, and for use with transitivity rules with more than 2
  8263 premises;
  8264 
  8265 * Isar/Pure: scalable support for case-analysis type proofs: new
  8266 'case' language element refers to local contexts symbolically, as
  8267 produced by certain proof methods; internally, case names are attached
  8268 to theorems as "tags";
  8269 
  8270 * Isar/Pure: theory command 'hide' removes declarations from
  8271 class/type/const name spaces;
  8272 
  8273 * Isar/Pure: theory command 'defs' supports option "(overloaded)" to
  8274 indicate potential overloading;
  8275 
  8276 * Isar/Pure: changed syntax of local blocks from {{ }} to { };
  8277 
  8278 * Isar/Pure: syntax of sorts made 'inner', i.e. have to write
  8279 "{a,b,c}" instead of {a,b,c};
  8280 
  8281 * Isar/Pure now provides its own version of intro/elim/dest
  8282 attributes; useful for building new logics, but beware of confusion
  8283 with the version in Provers/classical;
  8284 
  8285 * Isar/Pure: the local context of (non-atomic) goals is provided via
  8286 case name 'antecedent';
  8287 
  8288 * Isar/Pure: removed obsolete 'transfer' attribute (transfer of thms
  8289 to the current context is now done automatically);
  8290 
  8291 * Isar/Pure: theory command 'method_setup' provides a simple interface
  8292 for definining proof methods in ML;
  8293 
  8294 * Isar/Provers: intro/elim/dest attributes changed; renamed
  8295 intro/intro!/intro!! flags to intro!/intro/intro? (INCOMPATIBILITY, in
  8296 most cases, one should have to change intro!! to intro? only);
  8297 replaced "delrule" by "rule del";
  8298 
  8299 * Isar/Provers: new 'hypsubst' method, plain 'subst' method and
  8300 'symmetric' attribute (the latter supercedes [RS sym]);
  8301 
  8302 * Isar/Provers: splitter support (via 'split' attribute and 'simp'
  8303 method modifier); 'simp' method: 'only:' modifier removes loopers as
  8304 well (including splits);
  8305 
  8306 * Isar/Provers: Simplifier and Classical methods now support all kind
  8307 of modifiers used in the past, including 'cong', 'iff', etc.
  8308 
  8309 * Isar/Provers: added 'fastsimp' and 'clarsimp' methods (combination
  8310 of Simplifier and Classical reasoner);
  8311 
  8312 * Isar/HOL: new proof method 'cases' and improved version of 'induct'
  8313 now support named cases; major packages (inductive, datatype, primrec,
  8314 recdef) support case names and properly name parameters;
  8315 
  8316 * Isar/HOL: new transitivity rules for substitution in inequalities --
  8317 monotonicity conditions are extracted to be proven at end of
  8318 calculations;
  8319 
  8320 * Isar/HOL: removed 'case_split' thm binding, should use 'cases' proof
  8321 method anyway;
  8322 
  8323 * Isar/HOL: removed old expand_if = split_if; theorems if_splits =
  8324 split_if split_if_asm; datatype package provides theorems foo.splits =
  8325 foo.split foo.split_asm for each datatype;
  8326 
  8327 * Isar/HOL: tuned inductive package, rename "intrs" to "intros"
  8328 (potential INCOMPATIBILITY), emulation of mk_cases feature for proof
  8329 scripts: new 'inductive_cases' command and 'ind_cases' method; (Note:
  8330 use "(cases (simplified))" method in proper proof texts);
  8331 
  8332 * Isar/HOL: added global 'arith_split' attribute for 'arith' method;
  8333 
  8334 * Isar: names of theorems etc. may be natural numbers as well;
  8335 
  8336 * Isar: 'pr' command: optional arguments for goals_limit and
  8337 ProofContext.prems_limit; no longer prints theory contexts, but only
  8338 proof states;
  8339 
  8340 * Isar: diagnostic commands 'pr', 'thm', 'prop', 'term', 'typ' admit
  8341 additional print modes to be specified; e.g. "pr(latex)" will print
  8342 proof state according to the Isabelle LaTeX style;
  8343 
  8344 * Isar: improved support for emulating tactic scripts, including proof
  8345 methods 'rule_tac' etc., 'cut_tac', 'thin_tac', 'subgoal_tac',
  8346 'rename_tac', 'rotate_tac', 'tactic', and 'case_tac' / 'induct_tac'
  8347 (for HOL datatypes);
  8348 
  8349 * Isar: simplified (more robust) goal selection of proof methods: 1st
  8350 goal, all goals, or explicit goal specifier (tactic emulation); thus
  8351 'proof method scripts' have to be in depth-first order;
  8352 
  8353 * Isar: tuned 'let' syntax: replaced 'as' keyword by 'and';
  8354 
  8355 * Isar: removed 'help' command, which hasn't been too helpful anyway;
  8356 should instead use individual commands for printing items
  8357 (print_commands, print_methods etc.);
  8358 
  8359 * Isar: added 'nothing' --- the empty list of theorems;
  8360 
  8361 
  8362 *** HOL ***
  8363 
  8364 * HOL/MicroJava: formalization of a fragment of Java, together with a
  8365 corresponding virtual machine and a specification of its bytecode
  8366 verifier and a lightweight bytecode verifier, including proofs of
  8367 type-safety; by Gerwin Klein, Tobias Nipkow, David von Oheimb, and
  8368 Cornelia Pusch (see also the homepage of project Bali at
  8369 http://isabelle.in.tum.de/Bali/);
  8370 
  8371 * HOL/Algebra: new theory of rings and univariate polynomials, by
  8372 Clemens Ballarin;
  8373 
  8374 * HOL/NumberTheory: fundamental Theorem of Arithmetic, Chinese
  8375 Remainder Theorem, Fermat/Euler Theorem, Wilson's Theorem, by Thomas M
  8376 Rasmussen;
  8377 
  8378 * HOL/Lattice: fundamental concepts of lattice theory and order
  8379 structures, including duals, properties of bounds versus algebraic
  8380 laws, lattice operations versus set-theoretic ones, the Knaster-Tarski
  8381 Theorem for complete lattices etc.; may also serve as a demonstration
  8382 for abstract algebraic reasoning using axiomatic type classes, and
  8383 mathematics-style proof in Isabelle/Isar; by Markus Wenzel;
  8384 
  8385 * HOL/Prolog: a (bare-bones) implementation of Lambda-Prolog, by David
  8386 von Oheimb;
  8387 
  8388 * HOL/IMPP: extension of IMP with local variables and mutually
  8389 recursive procedures, by David von Oheimb;
  8390 
  8391 * HOL/Lambda: converted into new-style theory and document;
  8392 
  8393 * HOL/ex/Multiquote: example of multiple nested quotations and
  8394 anti-quotations -- basically a generalized version of de-Bruijn
  8395 representation; very useful in avoiding lifting of operations;
  8396 
  8397 * HOL/record: added general record equality rule to simpset; fixed
  8398 select-update simplification procedure to handle extended records as
  8399 well; admit "r" as field name;
  8400 
  8401 * HOL: 0 is now overloaded over the new sort "zero", allowing its use with
  8402 other numeric types and also as the identity of groups, rings, etc.;
  8403 
  8404 * HOL: new axclass plus_ac0 for addition with the AC-laws and 0 as identity.
  8405 Types nat and int belong to this axclass;
  8406 
  8407 * HOL: greatly improved simplification involving numerals of type nat, int, real:
  8408    (i + #8 + j) = Suc k simplifies to  #7 + (i + j) = k
  8409    i*j + k + j*#3*i     simplifies to  #4*(i*j) + k
  8410   two terms #m*u and #n*u are replaced by #(m+n)*u
  8411     (where #m, #n and u can implicitly be 1; this is simproc combine_numerals)
  8412   and the term/formula #m*u+x ~~ #n*u+y simplifies simplifies to #(m-n)+x ~~ y
  8413     or x ~~ #(n-m)+y, where ~~ is one of = < <= or - (simproc cancel_numerals);
  8414 
  8415 * HOL: meson_tac is available (previously in ex/meson.ML); it is a
  8416 powerful prover for predicate logic but knows nothing of clasets; see
  8417 ex/mesontest.ML and ex/mesontest2.ML for example applications;
  8418 
  8419 * HOL: new version of "case_tac" subsumes both boolean case split and
  8420 "exhaust_tac" on datatypes; INCOMPATIBILITY: exhaust_tac no longer
  8421 exists, may define val exhaust_tac = case_tac for ad-hoc portability;
  8422 
  8423 * HOL: simplification no longer dives into case-expressions: only the
  8424 selector expression is simplified, but not the remaining arms; to
  8425 enable full simplification of case-expressions for datatype t, you may
  8426 remove t.weak_case_cong from the simpset, either globally (Delcongs
  8427 [thm"t.weak_case_cong"];) or locally (delcongs [...]).
  8428 
  8429 * HOL/recdef: the recursion equations generated by 'recdef' for
  8430 function 'f' are now called f.simps instead of f.rules; if all
  8431 termination conditions are proved automatically, these simplification
  8432 rules are added to the simpset, as in primrec; rules may be named
  8433 individually as well, resulting in a separate list of theorems for
  8434 each equation;
  8435 
  8436 * HOL/While is a new theory that provides a while-combinator. It
  8437 permits the definition of tail-recursive functions without the
  8438 provision of a termination measure. The latter is necessary once the
  8439 invariant proof rule for while is applied.
  8440 
  8441 * HOL: new (overloaded) notation for the set of elements below/above
  8442 some element: {..u}, {..u(}, {l..}, {)l..}. See theory SetInterval.
  8443 
  8444 * HOL: theorems impI, allI, ballI bound as "strip";
  8445 
  8446 * HOL: new tactic induct_thm_tac: thm -> string -> int -> tactic
  8447 induct_tac th "x1 ... xn" expects th to have a conclusion of the form
  8448 P v1 ... vn and abbreviates res_inst_tac [("v1","x1"),...,("vn","xn")] th;
  8449 
  8450 * HOL/Real: "rabs" replaced by overloaded "abs" function;
  8451 
  8452 * HOL: theory Sexp now in HOL/Induct examples (it used to be part of
  8453 main HOL, but was unused);
  8454 
  8455 * HOL: fewer consts declared as global (e.g. have to refer to
  8456 "Lfp.lfp" instead of "lfp" internally; affects ML packages only);
  8457 
  8458 * HOL: tuned AST representation of nested pairs, avoiding bogus output
  8459 in case of overlap with user translations (e.g. judgements over
  8460 tuples); (note that the underlying logical represenation is still
  8461 bogus);
  8462 
  8463 
  8464 *** ZF ***
  8465 
  8466 * ZF: simplification automatically cancels common terms in arithmetic
  8467 expressions over nat and int;
  8468 
  8469 * ZF: new treatment of nat to minimize type-checking: all operators
  8470 coerce their operands to a natural number using the function natify,
  8471 making the algebraic laws unconditional;
  8472 
  8473 * ZF: as above, for int: operators coerce their operands to an integer
  8474 using the function intify;
  8475 
  8476 * ZF: the integer library now contains many of the usual laws for the
  8477 orderings, including $<=, and monotonicity laws for $+ and $*;
  8478 
  8479 * ZF: new example ZF/ex/NatSum to demonstrate integer arithmetic
  8480 simplification;
  8481 
  8482 * FOL and ZF: AddIffs now available, giving theorems of the form P<->Q
  8483 to the simplifier and classical reasoner simultaneously;
  8484 
  8485 
  8486 *** General ***
  8487 
  8488 * Provers: blast_tac now handles actual object-logic rules as
  8489 assumptions; note that auto_tac uses blast_tac internally as well;
  8490 
  8491 * Provers: new functions rulify/rulify_no_asm: thm -> thm for turning
  8492 outer -->/All/Ball into ==>/!!; qed_spec_mp now uses rulify_no_asm;
  8493 
  8494 * Provers: delrules now handles destruct rules as well (no longer need
  8495 explicit make_elim);
  8496 
  8497 * Provers: Blast_tac now warns of and ignores "weak elimination rules" e.g.
  8498   [| inj ?f;          ?f ?x = ?f ?y; ?x = ?y ==> ?W |] ==> ?W
  8499 use instead the strong form,
  8500   [| inj ?f; ~ ?W ==> ?f ?x = ?f ?y; ?x = ?y ==> ?W |] ==> ?W
  8501 in HOL, FOL and ZF the function cla_make_elim will create such rules
  8502 from destruct-rules;
  8503 
  8504 * Provers: Simplifier.easy_setup provides a fast path to basic
  8505 Simplifier setup for new object-logics;
  8506 
  8507 * Pure: AST translation rules no longer require constant head on LHS;
  8508 
  8509 * Pure: improved name spaces: ambiguous output is qualified; support
  8510 for hiding of names;
  8511 
  8512 * system: smart setup of canonical ML_HOME, ISABELLE_INTERFACE, and
  8513 XSYMBOL_HOME; no longer need to do manual configuration in most
  8514 situations;
  8515 
  8516 * system: compression of ML heaps images may now be controlled via -c
  8517 option of isabelle and isatool usedir (currently only observed by
  8518 Poly/ML);
  8519 
  8520 * system: isatool installfonts may handle X-Symbol fonts as well (very
  8521 useful for remote X11);
  8522 
  8523 * system: provide TAGS file for Isabelle sources;
  8524 
  8525 * ML: infix 'OF' is a version of 'MRS' with more appropriate argument
  8526 order;
  8527 
  8528 * ML: renamed flags Syntax.trace_norm_ast to Syntax.trace_ast; global
  8529 timing flag supersedes proof_timing and Toplevel.trace;
  8530 
  8531 * ML: new combinators |>> and |>>> for incremental transformations
  8532 with secondary results (e.g. certain theory extensions):
  8533 
  8534 * ML: PureThy.add_defs gets additional argument to indicate potential
  8535 overloading (usually false);
  8536 
  8537 * ML: PureThy.add_thms/add_axioms/add_defs now return theorems as
  8538 results;
  8539 
  8540 
  8541 
  8542 New in Isabelle99 (October 1999)
  8543 --------------------------------
  8544 
  8545 *** Overview of INCOMPATIBILITIES (see below for more details) ***
  8546 
  8547 * HOL: The THEN and ELSE parts of conditional expressions (if P then x else y)
  8548 are no longer simplified.  (This allows the simplifier to unfold recursive
  8549 functional programs.)  To restore the old behaviour, declare
  8550 
  8551     Delcongs [if_weak_cong];
  8552 
  8553 * HOL: Removed the obsolete syntax "Compl A"; use -A for set
  8554 complement;
  8555 
  8556 * HOL: the predicate "inj" is now defined by translation to "inj_on";
  8557 
  8558 * HOL/datatype: mutual_induct_tac no longer exists --
  8559   use induct_tac "x_1 ... x_n" instead of mutual_induct_tac ["x_1", ..., "x_n"]
  8560 
  8561 * HOL/typedef: fixed type inference for representing set; type
  8562 arguments now have to occur explicitly on the rhs as type constraints;
  8563 
  8564 * ZF: The con_defs part of an inductive definition may no longer refer
  8565 to constants declared in the same theory;
  8566 
  8567 * HOL, ZF: the function mk_cases, generated by the inductive
  8568 definition package, has lost an argument.  To simplify its result, it
  8569 uses the default simpset instead of a supplied list of theorems.
  8570 
  8571 * HOL/List: the constructors of type list are now Nil and Cons;
  8572 
  8573 * Simplifier: the type of the infix ML functions
  8574         setSSolver addSSolver setSolver addSolver
  8575 is now  simpset * solver -> simpset  where `solver' is a new abstract type
  8576 for packaging solvers. A solver is created via
  8577         mk_solver: string -> (thm list -> int -> tactic) -> solver
  8578 where the string argument is only a comment.
  8579 
  8580 
  8581 *** Proof tools ***
  8582 
  8583 * Provers/Arith/fast_lin_arith.ML contains a functor for creating a
  8584 decision procedure for linear arithmetic. Currently it is used for
  8585 types `nat', `int', and `real' in HOL (see below); it can, should and
  8586 will be instantiated for other types and logics as well.
  8587 
  8588 * The simplifier now accepts rewrite rules with flexible heads, eg
  8589      hom ?f ==> ?f(?x+?y) = ?f ?x + ?f ?y
  8590   They are applied like any rule with a non-pattern lhs, i.e. by first-order
  8591   matching.
  8592 
  8593 
  8594 *** General ***
  8595 
  8596 * New Isabelle/Isar subsystem provides an alternative to traditional
  8597 tactical theorem proving; together with the ProofGeneral/isar user
  8598 interface it offers an interactive environment for developing human
  8599 readable proof documents (Isar == Intelligible semi-automated
  8600 reasoning); for further information see isatool doc isar-ref,
  8601 src/HOL/Isar_examples and http://isabelle.in.tum.de/Isar/
  8602 
  8603 * improved and simplified presentation of theories: better HTML markup
  8604 (including colors), graph views in several sizes; isatool usedir now
  8605 provides a proper interface for user theories (via -P option); actual
  8606 document preparation based on (PDF)LaTeX is available as well (for
  8607 new-style theories only); see isatool doc system for more information;
  8608 
  8609 * native support for Proof General, both for classic Isabelle and
  8610 Isabelle/Isar;
  8611 
  8612 * ML function thm_deps visualizes dependencies of theorems and lemmas,
  8613 using the graph browser tool;
  8614 
  8615 * Isabelle manuals now also available as PDF;
  8616 
  8617 * theory loader rewritten from scratch (may not be fully
  8618 bug-compatible); old loadpath variable has been replaced by show_path,
  8619 add_path, del_path, reset_path functions; new operations such as
  8620 update_thy, touch_thy, remove_thy, use/update_thy_only (see also
  8621 isatool doc ref);
  8622 
  8623 * improved isatool install: option -k creates KDE application icon,
  8624 option -p DIR installs standalone binaries;
  8625 
  8626 * added ML_PLATFORM setting (useful for cross-platform installations);
  8627 more robust handling of platform specific ML images for SML/NJ;
  8628 
  8629 * the settings environment is now statically scoped, i.e. it is never
  8630 created again in sub-processes invoked from isabelle, isatool, or
  8631 Isabelle;
  8632 
  8633 * path element specification '~~' refers to '$ISABELLE_HOME';
  8634 
  8635 * in locales, the "assumes" and "defines" parts may be omitted if
  8636 empty;
  8637 
  8638 * new print_mode "xsymbols" for extended symbol support (e.g. genuine
  8639 long arrows);
  8640 
  8641 * new print_mode "HTML";
  8642 
  8643 * new flag show_tags controls display of tags of theorems (which are
  8644 basically just comments that may be attached by some tools);
  8645 
  8646 * Isamode 2.6 requires patch to accomodate change of Isabelle font
  8647 mode and goal output format:
  8648 
  8649 diff -r Isamode-2.6/elisp/isa-load.el Isamode/elisp/isa-load.el
  8650 244c244
  8651 <       (list (isa-getenv "ISABELLE") "-msymbols" logic-name)
  8652 ---
  8653 >       (list (isa-getenv "ISABELLE") "-misabelle_font" "-msymbols" logic-name)
  8654 diff -r Isabelle-2.6/elisp/isa-proofstate.el Isamode/elisp/isa-proofstate.el
  8655 181c181
  8656 < (defconst proofstate-proofstart-regexp "^Level [0-9]+$"
  8657 ---
  8658 > (defconst proofstate-proofstart-regexp "^Level [0-9]+"
  8659 
  8660 * function bind_thms stores lists of theorems (cf. bind_thm);
  8661 
  8662 * new shorthand tactics ftac, eatac, datac, fatac;
  8663 
  8664 * qed (and friends) now accept "" as result name; in that case the
  8665 theorem is not stored, but proper checks and presentation of the
  8666 result still apply;
  8667 
  8668 * theorem database now also indexes constants "Trueprop", "all",
  8669 "==>", "=="; thus thms_containing, findI etc. may retrieve more rules;
  8670 
  8671 
  8672 *** HOL ***
  8673 
  8674 ** HOL arithmetic **
  8675 
  8676 * There are now decision procedures for linear arithmetic over nat and
  8677 int:
  8678 
  8679 1. arith_tac copes with arbitrary formulae involving `=', `<', `<=',
  8680 `+', `-', `Suc', `min', `max' and numerical constants; other subterms
  8681 are treated as atomic; subformulae not involving type `nat' or `int'
  8682 are ignored; quantified subformulae are ignored unless they are
  8683 positive universal or negative existential. The tactic has to be
  8684 invoked by hand and can be a little bit slow. In particular, the
  8685 running time is exponential in the number of occurrences of `min' and
  8686 `max', and `-' on `nat'.
  8687 
  8688 2. fast_arith_tac is a cut-down version of arith_tac: it only takes
  8689 (negated) (in)equalities among the premises and the conclusion into
  8690 account (i.e. no compound formulae) and does not know about `min' and
  8691 `max', and `-' on `nat'. It is fast and is used automatically by the
  8692 simplifier.
  8693 
  8694 NB: At the moment, these decision procedures do not cope with mixed
  8695 nat/int formulae where the two parts interact, such as `m < n ==>
  8696 int(m) < int(n)'.
  8697 
  8698 * HOL/Numeral provides a generic theory of numerals (encoded
  8699 efficiently as bit strings); setup for types nat/int/real is in place;
  8700 INCOMPATIBILITY: since numeral syntax is now polymorphic, rather than
  8701 int, existing theories and proof scripts may require a few additional
  8702 type constraints;
  8703 
  8704 * integer division and remainder can now be performed on constant
  8705 arguments;
  8706 
  8707 * many properties of integer multiplication, division and remainder
  8708 are now available;
  8709 
  8710 * An interface to the Stanford Validity Checker (SVC) is available through the
  8711 tactic svc_tac.  Propositional tautologies and theorems of linear arithmetic
  8712 are proved automatically.  SVC must be installed separately, and its results
  8713 must be TAKEN ON TRUST (Isabelle does not check the proofs, but tags any
  8714 invocation of the underlying oracle).  For SVC see
  8715   http://verify.stanford.edu/SVC
  8716 
  8717 * IsaMakefile: the HOL-Real target now builds an actual image;
  8718 
  8719 
  8720 ** HOL misc **
  8721 
  8722 * HOL/Real/HahnBanach: the Hahn-Banach theorem for real vector spaces
  8723 (in Isabelle/Isar) -- by Gertrud Bauer;
  8724 
  8725 * HOL/BCV: generic model of bytecode verification, i.e. data-flow
  8726 analysis for assembly languages with subtypes;
  8727 
  8728 * HOL/TLA (Lamport's Temporal Logic of Actions): major reorganization
  8729 -- avoids syntactic ambiguities and treats state, transition, and
  8730 temporal levels more uniformly; introduces INCOMPATIBILITIES due to
  8731 changed syntax and (many) tactics;
  8732 
  8733 * HOL/inductive: Now also handles more general introduction rules such
  8734   as "ALL y. (y, x) : r --> y : acc r ==> x : acc r"; monotonicity
  8735   theorems are now maintained within the theory (maintained via the
  8736   "mono" attribute);
  8737 
  8738 * HOL/datatype: Now also handles arbitrarily branching datatypes
  8739   (using function types) such as
  8740 
  8741   datatype 'a tree = Atom 'a | Branch "nat => 'a tree"
  8742 
  8743 * HOL/record: record_simproc (part of the default simpset) takes care
  8744 of selectors applied to updated records; record_split_tac is no longer
  8745 part of the default claset; update_defs may now be removed from the
  8746 simpset in many cases; COMPATIBILITY: old behavior achieved by
  8747 
  8748   claset_ref () := claset() addSWrapper record_split_wrapper;
  8749   Delsimprocs [record_simproc]
  8750 
  8751 * HOL/typedef: fixed type inference for representing set; type
  8752 arguments now have to occur explicitly on the rhs as type constraints;
  8753 
  8754 * HOL/recdef (TFL): 'congs' syntax now expects comma separated list of theorem
  8755 names rather than an ML expression;
  8756 
  8757 * HOL/defer_recdef (TFL): like recdef but the well-founded relation can be
  8758 supplied later.  Program schemes can be defined, such as
  8759     "While B C s = (if B s then While B C (C s) else s)"
  8760 where the well-founded relation can be chosen after B and C have been given.
  8761 
  8762 * HOL/List: the constructors of type list are now Nil and Cons;
  8763 INCOMPATIBILITY: while [] and infix # syntax is still there, of
  8764 course, ML tools referring to List.list.op # etc. have to be adapted;
  8765 
  8766 * HOL_quantifiers flag superseded by "HOL" print mode, which is
  8767 disabled by default; run isabelle with option -m HOL to get back to
  8768 the original Gordon/HOL-style output;
  8769 
  8770 * HOL/Ord.thy: new bounded quantifier syntax (input only): ALL x<y. P,
  8771 ALL x<=y. P, EX x<y. P, EX x<=y. P;
  8772 
  8773 * HOL basic syntax simplified (more orthogonal): all variants of
  8774 All/Ex now support plain / symbolic / HOL notation; plain syntax for
  8775 Eps operator is provided as well: "SOME x. P[x]";
  8776 
  8777 * HOL/Sum.thy: sum_case has been moved to HOL/Datatype;
  8778 
  8779 * HOL/Univ.thy: infix syntax <*>, <+>, <**>, <+> eliminated and made
  8780 thus available for user theories;
  8781 
  8782 * HOLCF/IOA/Sequents: renamed 'Cons' to 'Consq' to avoid clash with
  8783 HOL/List; hardly an INCOMPATIBILITY since '>>' syntax is used all the
  8784 time;
  8785 
  8786 * HOL: new tactic smp_tac: int -> int -> tactic, which applies spec
  8787 several times and then mp;
  8788 
  8789 
  8790 *** LK ***
  8791 
  8792 * the notation <<...>> is now available as a notation for sequences of
  8793 formulas;
  8794 
  8795 * the simplifier is now installed
  8796 
  8797 * the axiom system has been generalized (thanks to Soren Heilmann)
  8798 
  8799 * the classical reasoner now has a default rule database
  8800 
  8801 
  8802 *** ZF ***
  8803 
  8804 * new primrec section allows primitive recursive functions to be given
  8805 directly (as in HOL) over datatypes and the natural numbers;
  8806 
  8807 * new tactics induct_tac and exhaust_tac for induction (or case
  8808 analysis) over datatypes and the natural numbers;
  8809 
  8810 * the datatype declaration of type T now defines the recursor T_rec;
  8811 
  8812 * simplification automatically does freeness reasoning for datatype
  8813 constructors;
  8814 
  8815 * automatic type-inference, with AddTCs command to insert new
  8816 type-checking rules;
  8817 
  8818 * datatype introduction rules are now added as Safe Introduction rules
  8819 to the claset;
  8820 
  8821 * the syntax "if P then x else y" is now available in addition to
  8822 if(P,x,y);
  8823 
  8824 
  8825 *** Internal programming interfaces ***
  8826 
  8827 * tuned simplifier trace output; new flag debug_simp;
  8828 
  8829 * structures Vartab / Termtab (instances of TableFun) offer efficient
  8830 tables indexed by indexname_ord / term_ord (compatible with aconv);
  8831 
  8832 * AxClass.axclass_tac lost the theory argument;
  8833 
  8834 * tuned current_goals_markers semantics: begin / end goal avoids
  8835 printing empty lines;
  8836 
  8837 * removed prs and prs_fn hook, which was broken because it did not
  8838 include \n in its semantics, forcing writeln to add one
  8839 uncoditionally; replaced prs_fn by writeln_fn; consider std_output:
  8840 string -> unit if you really want to output text without newline;
  8841 
  8842 * Symbol.output subject to print mode; INCOMPATIBILITY: defaults to
  8843 plain output, interface builders may have to enable 'isabelle_font'
  8844 mode to get Isabelle font glyphs as before;
  8845 
  8846 * refined token_translation interface; INCOMPATIBILITY: output length
  8847 now of type real instead of int;
  8848 
  8849 * theory loader actions may be traced via new ThyInfo.add_hook
  8850 interface (see src/Pure/Thy/thy_info.ML); example application: keep
  8851 your own database of information attached to *whole* theories -- as
  8852 opposed to intra-theory data slots offered via TheoryDataFun;
  8853 
  8854 * proper handling of dangling sort hypotheses (at last!);
  8855 Thm.strip_shyps and Drule.strip_shyps_warning take care of removing
  8856 extra sort hypotheses that can be witnessed from the type signature;
  8857 the force_strip_shyps flag is gone, any remaining shyps are simply
  8858 left in the theorem (with a warning issued by strip_shyps_warning);
  8859 
  8860 
  8861 
  8862 New in Isabelle98-1 (October 1998)
  8863 ----------------------------------
  8864 
  8865 *** Overview of INCOMPATIBILITIES (see below for more details) ***
  8866 
  8867 * several changes of automated proof tools;
  8868 
  8869 * HOL: major changes to the inductive and datatype packages, including
  8870 some minor incompatibilities of theory syntax;
  8871 
  8872 * HOL: renamed r^-1 to 'converse' from 'inverse'; 'inj_onto' is now
  8873 called `inj_on';
  8874 
  8875 * HOL: removed duplicate thms in Arith:
  8876   less_imp_add_less  should be replaced by  trans_less_add1
  8877   le_imp_add_le      should be replaced by  trans_le_add1
  8878 
  8879 * HOL: unary minus is now overloaded (new type constraints may be
  8880 required);
  8881 
  8882 * HOL and ZF: unary minus for integers is now #- instead of #~.  In
  8883 ZF, expressions such as n#-1 must be changed to n#- 1, since #-1 is
  8884 now taken as an integer constant.
  8885 
  8886 * Pure: ML function 'theory_of' renamed to 'theory';
  8887 
  8888 
  8889 *** Proof tools ***
  8890 
  8891 * Simplifier:
  8892   1. Asm_full_simp_tac is now more aggressive.
  8893      1. It will sometimes reorient premises if that increases their power to
  8894         simplify.
  8895      2. It does no longer proceed strictly from left to right but may also
  8896         rotate premises to achieve further simplification.
  8897      For compatibility reasons there is now Asm_lr_simp_tac which is like the
  8898      old Asm_full_simp_tac in that it does not rotate premises.
  8899   2. The simplifier now knows a little bit about nat-arithmetic.
  8900 
  8901 * Classical reasoner: wrapper mechanism for the classical reasoner now
  8902 allows for selected deletion of wrappers, by introduction of names for
  8903 wrapper functionals.  This implies that addbefore, addSbefore,
  8904 addaltern, and addSaltern now take a pair (name, tactic) as argument,
  8905 and that adding two tactics with the same name overwrites the first
  8906 one (emitting a warning).
  8907   type wrapper = (int -> tactic) -> (int -> tactic)
  8908   setWrapper, setSWrapper, compWrapper and compSWrapper are replaced by
  8909   addWrapper, addSWrapper: claset * (string * wrapper) -> claset
  8910   delWrapper, delSWrapper: claset *  string            -> claset
  8911   getWrapper is renamed to appWrappers, getSWrapper to appSWrappers;
  8912 
  8913 * Classical reasoner: addbefore/addSbefore now have APPEND/ORELSE
  8914 semantics; addbefore now affects only the unsafe part of step_tac
  8915 etc.; this affects addss/auto_tac/force_tac, so EXISTING PROOFS MAY
  8916 FAIL, but proofs should be fixable easily, e.g. by replacing Auto_tac
  8917 by Force_tac;
  8918 
  8919 * Classical reasoner: setwrapper to setWrapper and compwrapper to
  8920 compWrapper; added safe wrapper (and access functions for it);
  8921 
  8922 * HOL/split_all_tac is now much faster and fails if there is nothing
  8923 to split.  Some EXISTING PROOFS MAY REQUIRE ADAPTION because the order
  8924 and the names of the automatically generated variables have changed.
  8925 split_all_tac has moved within claset() from unsafe wrappers to safe
  8926 wrappers, which means that !!-bound variables are split much more
  8927 aggressively, and safe_tac and clarify_tac now split such variables.
  8928 If this splitting is not appropriate, use delSWrapper "split_all_tac".
  8929 Note: the same holds for record_split_tac, which does the job of
  8930 split_all_tac for record fields.
  8931 
  8932 * HOL/Simplifier: Rewrite rules for case distinctions can now be added
  8933 permanently to the default simpset using Addsplits just like
  8934 Addsimps. They can be removed via Delsplits just like
  8935 Delsimps. Lower-case versions are also available.
  8936 
  8937 * HOL/Simplifier: The rule split_if is now part of the default
  8938 simpset. This means that the simplifier will eliminate all occurrences
  8939 of if-then-else in the conclusion of a goal. To prevent this, you can
  8940 either remove split_if completely from the default simpset by
  8941 `Delsplits [split_if]' or remove it in a specific call of the
  8942 simplifier using `... delsplits [split_if]'.  You can also add/delete
  8943 other case splitting rules to/from the default simpset: every datatype
  8944 generates suitable rules `split_t_case' and `split_t_case_asm' (where
  8945 t is the name of the datatype).
  8946 
  8947 * Classical reasoner / Simplifier combination: new force_tac (and
  8948 derivatives Force_tac, force) combines rewriting and classical
  8949 reasoning (and whatever other tools) similarly to auto_tac, but is
  8950 aimed to solve the given subgoal completely.
  8951 
  8952 
  8953 *** General ***
  8954 
  8955 * new top-level commands `Goal' and `Goalw' that improve upon `goal'
  8956 and `goalw': the theory is no longer needed as an explicit argument -
  8957 the current theory context is used; assumptions are no longer returned
  8958 at the ML-level unless one of them starts with ==> or !!; it is
  8959 recommended to convert to these new commands using isatool fixgoal
  8960 (backup your sources first!);
  8961 
  8962 * new top-level commands 'thm' and 'thms' for retrieving theorems from
  8963 the current theory context, and 'theory' to lookup stored theories;
  8964 
  8965 * new theory section 'locale' for declaring constants, assumptions and
  8966 definitions that have local scope;
  8967 
  8968 * new theory section 'nonterminals' for purely syntactic types;
  8969 
  8970 * new theory section 'setup' for generic ML setup functions
  8971 (e.g. package initialization);
  8972 
  8973 * the distribution now includes Isabelle icons: see
  8974 lib/logo/isabelle-{small,tiny}.xpm;
  8975 
  8976 * isatool install - install binaries with absolute references to
  8977 ISABELLE_HOME/bin;
  8978 
  8979 * isatool logo -- create instances of the Isabelle logo (as EPS);
  8980 
  8981 * print mode 'emacs' reserved for Isamode;
  8982 
  8983 * support multiple print (ast) translations per constant name;
  8984 
  8985 * theorems involving oracles are now printed with a suffixed [!];
  8986 
  8987 
  8988 *** HOL ***
  8989 
  8990 * there is now a tutorial on Isabelle/HOL (do 'isatool doc tutorial');
  8991 
  8992 * HOL/inductive package reorganized and improved: now supports mutual
  8993 definitions such as
  8994 
  8995   inductive EVEN ODD
  8996     intrs
  8997       null "0 : EVEN"
  8998       oddI "n : EVEN ==> Suc n : ODD"
  8999       evenI "n : ODD ==> Suc n : EVEN"
  9000 
  9001 new theorem list "elims" contains an elimination rule for each of the
  9002 recursive sets; inductive definitions now handle disjunctive premises
  9003 correctly (also ZF);
  9004 
  9005 INCOMPATIBILITIES: requires Inductive as an ancestor; component
  9006 "mutual_induct" no longer exists - the induction rule is always
  9007 contained in "induct";
  9008 
  9009 
  9010 * HOL/datatype package re-implemented and greatly improved: now
  9011 supports mutually recursive datatypes such as
  9012 
  9013   datatype
  9014     'a aexp = IF_THEN_ELSE ('a bexp) ('a aexp) ('a aexp)
  9015             | SUM ('a aexp) ('a aexp)
  9016             | DIFF ('a aexp) ('a aexp)
  9017             | NUM 'a
  9018   and
  9019     'a bexp = LESS ('a aexp) ('a aexp)
  9020             | AND ('a bexp) ('a bexp)
  9021             | OR ('a bexp) ('a bexp)
  9022 
  9023 as well as indirectly recursive datatypes such as
  9024 
  9025   datatype
  9026     ('a, 'b) term = Var 'a
  9027                   | App 'b ((('a, 'b) term) list)
  9028 
  9029 The new tactic  mutual_induct_tac [<var_1>, ..., <var_n>] i  performs
  9030 induction on mutually / indirectly recursive datatypes.
  9031 
  9032 Primrec equations are now stored in theory and can be accessed via
  9033 <function_name>.simps.
  9034 
  9035 INCOMPATIBILITIES:
  9036 
  9037   - Theories using datatypes must now have theory Datatype as an
  9038     ancestor.
  9039   - The specific <typename>.induct_tac no longer exists - use the
  9040     generic induct_tac instead.
  9041   - natE has been renamed to nat.exhaust - use exhaust_tac
  9042     instead of res_inst_tac ... natE. Note that the variable
  9043     names in nat.exhaust differ from the names in natE, this
  9044     may cause some "fragile" proofs to fail.
  9045   - The theorems split_<typename>_case and split_<typename>_case_asm
  9046     have been renamed to <typename>.split and <typename>.split_asm.
  9047   - Since default sorts of type variables are now handled correctly,
  9048     some datatype definitions may have to be annotated with explicit
  9049     sort constraints.
  9050   - Primrec definitions no longer require function name and type
  9051     of recursive argument.
  9052 
  9053 Consider using isatool fixdatatype to adapt your theories and proof
  9054 scripts to the new package (backup your sources first!).
  9055 
  9056 
  9057 * HOL/record package: considerably improved implementation; now
  9058 includes concrete syntax for record types, terms, updates; theorems
  9059 for surjective pairing and splitting !!-bound record variables; proof
  9060 support is as follows:
  9061 
  9062   1) standard conversions (selectors or updates applied to record
  9063 constructor terms) are part of the standard simpset;
  9064 
  9065   2) inject equations of the form ((x, y) = (x', y')) == x=x' & y=y' are
  9066 made part of standard simpset and claset via addIffs;
  9067 
  9068   3) a tactic for record field splitting (record_split_tac) is part of
  9069 the standard claset (addSWrapper);
  9070 
  9071 To get a better idea about these rules you may retrieve them via
  9072 something like 'thms "foo.simps"' or 'thms "foo.iffs"', where "foo" is
  9073 the name of your record type.
  9074 
  9075 The split tactic 3) conceptually simplifies by the following rule:
  9076 
  9077   "(!!x. PROP ?P x) == (!!a b. PROP ?P (a, b))"
  9078 
  9079 Thus any record variable that is bound by meta-all will automatically
  9080 blow up into some record constructor term, consequently the
  9081 simplifications of 1), 2) apply.  Thus force_tac, auto_tac etc. shall
  9082 solve record problems automatically.
  9083 
  9084 
  9085 * reorganized the main HOL image: HOL/Integ and String loaded by
  9086 default; theory Main includes everything;
  9087 
  9088 * automatic simplification of integer sums and comparisons, using cancellation;
  9089 
  9090 * added option_map_eq_Some and not_Some_eq to the default simpset and claset;
  9091 
  9092 * added disj_not1 = "(~P | Q) = (P --> Q)" to the default simpset;
  9093 
  9094 * many new identities for unions, intersections, set difference, etc.;
  9095 
  9096 * expand_if, expand_split, expand_sum_case and expand_nat_case are now
  9097 called split_if, split_split, split_sum_case and split_nat_case (to go
  9098 with add/delsplits);
  9099 
  9100 * HOL/Prod introduces simplification procedure unit_eq_proc rewriting
  9101 (?x::unit) = (); this is made part of the default simpset, which COULD
  9102 MAKE EXISTING PROOFS FAIL under rare circumstances (consider
  9103 'Delsimprocs [unit_eq_proc];' as last resort); also note that
  9104 unit_abs_eta_conv is added in order to counter the effect of
  9105 unit_eq_proc on (%u::unit. f u), replacing it by f rather than by
  9106 %u.f();
  9107 
  9108 * HOL/Fun INCOMPATIBILITY: `inj_onto' is now called `inj_on' (which
  9109 makes more sense);
  9110 
  9111 * HOL/Set INCOMPATIBILITY: rule `equals0D' is now a well-formed destruct rule;
  9112   It and 'sym RS equals0D' are now in the default  claset, giving automatic
  9113   disjointness reasoning but breaking a few old proofs.
  9114 
  9115 * HOL/Relation INCOMPATIBILITY: renamed the relational operator r^-1
  9116 to 'converse' from 'inverse' (for compatibility with ZF and some
  9117 literature);
  9118 
  9119 * HOL/recdef can now declare non-recursive functions, with {} supplied as
  9120 the well-founded relation;
  9121 
  9122 * HOL/Set INCOMPATIBILITY: the complement of set A is now written -A instead of
  9123     Compl A.  The "Compl" syntax remains available as input syntax for this
  9124     release ONLY.
  9125 
  9126 * HOL/Update: new theory of function updates:
  9127     f(a:=b) == %x. if x=a then b else f x
  9128 may also be iterated as in f(a:=b,c:=d,...);
  9129 
  9130 * HOL/Vimage: new theory for inverse image of a function, syntax f-``B;
  9131 
  9132 * HOL/List:
  9133   - new function list_update written xs[i:=v] that updates the i-th
  9134     list position. May also be iterated as in xs[i:=a,j:=b,...].
  9135   - new function `upt' written [i..j(] which generates the list
  9136     [i,i+1,...,j-1], i.e. the upper bound is excluded. To include the upper
  9137     bound write [i..j], which is a shorthand for [i..j+1(].
  9138   - new lexicographic orderings and corresponding wellfoundedness theorems.
  9139 
  9140 * HOL/Arith:
  9141   - removed 'pred' (predecessor) function;
  9142   - generalized some theorems about n-1;
  9143   - many new laws about "div" and "mod";
  9144   - new laws about greatest common divisors (see theory ex/Primes);
  9145 
  9146 * HOL/Relation: renamed the relational operator r^-1 "converse"
  9147 instead of "inverse";
  9148 
  9149 * HOL/Induct/Multiset: a theory of multisets, including the wellfoundedness
  9150   of the multiset ordering;
  9151 
  9152 * directory HOL/Real: a construction of the reals using Dedekind cuts
  9153   (not included by default);
  9154 
  9155 * directory HOL/UNITY: Chandy and Misra's UNITY formalism;
  9156 
  9157 * directory HOL/Hoare: a new version of Hoare logic which permits many-sorted
  9158   programs, i.e. different program variables may have different types.
  9159 
  9160 * calling (stac rew i) now fails if "rew" has no effect on the goal
  9161   [previously, this check worked only if the rewrite rule was unconditional]
  9162   Now rew can involve either definitions or equalities (either == or =).
  9163 
  9164 
  9165 *** ZF ***
  9166 
  9167 * theory Main includes everything; INCOMPATIBILITY: theory ZF.thy contains
  9168   only the theorems proved on ZF.ML;
  9169 
  9170 * ZF INCOMPATIBILITY: rule `equals0D' is now a well-formed destruct rule;
  9171   It and 'sym RS equals0D' are now in the default  claset, giving automatic
  9172   disjointness reasoning but breaking a few old proofs.
  9173 
  9174 * ZF/Update: new theory of function updates
  9175     with default rewrite rule  f(x:=y) ` z = if(z=x, y, f`z)
  9176   may also be iterated as in f(a:=b,c:=d,...);
  9177 
  9178 * in  let x=t in u(x), neither t nor u(x) has to be an FOL term.
  9179 
  9180 * calling (stac rew i) now fails if "rew" has no effect on the goal
  9181   [previously, this check worked only if the rewrite rule was unconditional]
  9182   Now rew can involve either definitions or equalities (either == or =).
  9183 
  9184 * case_tac provided for compatibility with HOL
  9185     (like the old excluded_middle_tac, but with subgoals swapped)
  9186 
  9187 
  9188 *** Internal programming interfaces ***
  9189 
  9190 * Pure: several new basic modules made available for general use, see
  9191 also src/Pure/README;
  9192 
  9193 * improved the theory data mechanism to support encapsulation (data
  9194 kind name replaced by private Object.kind, acting as authorization
  9195 key); new type-safe user interface via functor TheoryDataFun; generic
  9196 print_data function becomes basically useless;
  9197 
  9198 * removed global_names compatibility flag -- all theory declarations
  9199 are qualified by default;
  9200 
  9201 * module Pure/Syntax now offers quote / antiquote translation
  9202 functions (useful for Hoare logic etc. with implicit dependencies);
  9203 see HOL/ex/Antiquote for an example use;
  9204 
  9205 * Simplifier now offers conversions (asm_)(full_)rewrite: simpset ->
  9206 cterm -> thm;
  9207 
  9208 * new tactical CHANGED_GOAL for checking that a tactic modifies a
  9209 subgoal;
  9210 
  9211 * Display.print_goals function moved to Locale.print_goals;
  9212 
  9213 * standard print function for goals supports current_goals_markers
  9214 variable for marking begin of proof, end of proof, start of goal; the
  9215 default is ("", "", ""); setting current_goals_markers := ("<proof>",
  9216 "</proof>", "<goal>") causes SGML like tagged proof state printing,
  9217 for example;
  9218 
  9219 
  9220 
  9221 New in Isabelle98 (January 1998)
  9222 --------------------------------
  9223 
  9224 *** Overview of INCOMPATIBILITIES (see below for more details) ***
  9225 
  9226 * changed lexical syntax of terms / types: dots made part of long
  9227 identifiers, e.g. "%x.x" no longer possible, should be "%x. x";
  9228 
  9229 * simpset (and claset) reference variable replaced by functions
  9230 simpset / simpset_ref;
  9231 
  9232 * no longer supports theory aliases (via merge) and non-trivial
  9233 implicit merge of thms' signatures;
  9234 
  9235 * most internal names of constants changed due to qualified names;
  9236 
  9237 * changed Pure/Sequence interface (see Pure/seq.ML);
  9238 
  9239 
  9240 *** General Changes ***
  9241 
  9242 * hierachically structured name spaces (for consts, types, axms, thms
  9243 etc.); new lexical class 'longid' (e.g. Foo.bar.x) may render much of
  9244 old input syntactically incorrect (e.g. "%x.x"); COMPATIBILITY:
  9245 isatool fixdots ensures space after dots (e.g. "%x. x"); set
  9246 long_names for fully qualified output names; NOTE: ML programs
  9247 (special tactics, packages etc.) referring to internal names may have
  9248 to be adapted to cope with fully qualified names; in case of severe
  9249 backward campatibility problems try setting 'global_names' at compile
  9250 time to have enrything declared within a flat name space; one may also
  9251 fine tune name declarations in theories via the 'global' and 'local'
  9252 section;
  9253 
  9254 * reimplemented the implicit simpset and claset using the new anytype
  9255 data filed in signatures; references simpset:simpset ref etc. are
  9256 replaced by functions simpset:unit->simpset and
  9257 simpset_ref:unit->simpset ref; COMPATIBILITY: use isatool fixclasimp
  9258 to patch your ML files accordingly;
  9259 
  9260 * HTML output now includes theory graph data for display with Java
  9261 applet or isatool browser; data generated automatically via isatool
  9262 usedir (see -i option, ISABELLE_USEDIR_OPTIONS);
  9263 
  9264 * defs may now be conditional; improved rewrite_goals_tac to handle
  9265 conditional equations;
  9266 
  9267 * defs now admits additional type arguments, using TYPE('a) syntax;
  9268 
  9269 * theory aliases via merge (e.g. M=A+B+C) no longer supported, always
  9270 creates a new theory node; implicit merge of thms' signatures is
  9271 restricted to 'trivial' ones; COMPATIBILITY: one may have to use
  9272 transfer:theory->thm->thm in (rare) cases;
  9273 
  9274 * improved handling of draft signatures / theories; draft thms (and
  9275 ctyps, cterms) are automatically promoted to real ones;
  9276 
  9277 * slightly changed interfaces for oracles: admit many per theory, named
  9278 (e.g. oracle foo = mlfun), additional name argument for invoke_oracle;
  9279 
  9280 * print_goals: optional output of const types (set show_consts and
  9281 show_types);
  9282 
  9283 * improved output of warnings (###) and errors (***);
  9284 
  9285 * subgoal_tac displays a warning if the new subgoal has type variables;
  9286 
  9287 * removed old README and Makefiles;
  9288 
  9289 * replaced print_goals_ref hook by print_current_goals_fn and result_error_fn;
  9290 
  9291 * removed obsolete init_pps and init_database;
  9292 
  9293 * deleted the obsolete tactical STATE, which was declared by
  9294     fun STATE tacfun st = tacfun st st;
  9295 
  9296 * cd and use now support path variables, e.g. $ISABELLE_HOME, or ~
  9297 (which abbreviates $HOME);
  9298 
  9299 * changed Pure/Sequence interface (see Pure/seq.ML); COMPATIBILITY:
  9300 use isatool fixseq to adapt your ML programs (this works for fully
  9301 qualified references to the Sequence structure only!);
  9302 
  9303 * use_thy no longer requires writable current directory; it always
  9304 reloads .ML *and* .thy file, if either one is out of date;
  9305 
  9306 
  9307 *** Classical Reasoner ***
  9308 
  9309 * Clarify_tac, clarify_tac, clarify_step_tac, Clarify_step_tac: new
  9310 tactics that use classical reasoning to simplify a subgoal without
  9311 splitting it into several subgoals;
  9312 
  9313 * Safe_tac: like safe_tac but uses the default claset;
  9314 
  9315 
  9316 *** Simplifier ***
  9317 
  9318 * added simplification meta rules:
  9319     (asm_)(full_)simplify: simpset -> thm -> thm;
  9320 
  9321 * simplifier.ML no longer part of Pure -- has to be loaded by object
  9322 logics (again);
  9323 
  9324 * added prems argument to simplification procedures;
  9325 
  9326 * HOL, FOL, ZF: added infix function `addsplits':
  9327   instead of `<simpset> setloop (split_tac <thms>)'
  9328   you can simply write `<simpset> addsplits <thms>'
  9329 
  9330 
  9331 *** Syntax ***
  9332 
  9333 * TYPE('a) syntax for type reflection terms;
  9334 
  9335 * no longer handles consts with name "" -- declare as 'syntax' instead;
  9336 
  9337 * pretty printer: changed order of mixfix annotation preference (again!);
  9338 
  9339 * Pure: fixed idt/idts vs. pttrn/pttrns syntactic categories;
  9340 
  9341 
  9342 *** HOL ***
  9343 
  9344 * HOL: there is a new splitter `split_asm_tac' that can be used e.g.
  9345   with `addloop' of the simplifier to faciliate case splitting in premises.
  9346 
  9347 * HOL/TLA: Stephan Merz's formalization of Lamport's Temporal Logic of Actions;
  9348 
  9349 * HOL/Auth: new protocol proofs including some for the Internet
  9350   protocol TLS;
  9351 
  9352 * HOL/Map: new theory of `maps' a la VDM;
  9353 
  9354 * HOL/simplifier: simplification procedures nat_cancel_sums for
  9355 cancelling out common nat summands from =, <, <= (in)equalities, or
  9356 differences; simplification procedures nat_cancel_factor for
  9357 cancelling common factor from =, <, <= (in)equalities over natural
  9358 sums; nat_cancel contains both kinds of procedures, it is installed by
  9359 default in Arith.thy -- this COULD MAKE EXISTING PROOFS FAIL;
  9360 
  9361 * HOL/simplifier: terms of the form
  9362   `? x. P1(x) & ... & Pn(x) & x=t & Q1(x) & ... Qn(x)'  (or t=x)
  9363   are rewritten to
  9364   `P1(t) & ... & Pn(t) & Q1(t) & ... Qn(t)',
  9365   and those of the form
  9366   `! x. P1(x) & ... & Pn(x) & x=t & Q1(x) & ... Qn(x) --> R(x)'  (or t=x)
  9367   are rewritten to
  9368   `P1(t) & ... & Pn(t) & Q1(t) & ... Qn(t) --> R(t)',
  9369 
  9370 * HOL/datatype
  9371   Each datatype `t' now comes with a theorem `split_t_case' of the form
  9372 
  9373   P(t_case f1 ... fn x) =
  9374      ( (!y1 ... ym1. x = C1 y1 ... ym1 --> P(f1 y1 ... ym1)) &
  9375         ...
  9376        (!y1 ... ymn. x = Cn y1 ... ymn --> P(f1 y1 ... ymn))
  9377      )
  9378 
  9379   and a theorem `split_t_case_asm' of the form
  9380 
  9381   P(t_case f1 ... fn x) =
  9382     ~( (? y1 ... ym1. x = C1 y1 ... ym1 & ~P(f1 y1 ... ym1)) |
  9383         ...
  9384        (? y1 ... ymn. x = Cn y1 ... ymn & ~P(f1 y1 ... ymn))
  9385      )
  9386   which can be added to a simpset via `addsplits'. The existing theorems
  9387   expand_list_case and expand_option_case have been renamed to
  9388   split_list_case and split_option_case.
  9389 
  9390 * HOL/Arithmetic:
  9391   - `pred n' is automatically converted to `n-1'.
  9392     Users are strongly encouraged not to use `pred' any longer,
  9393     because it will disappear altogether at some point.
  9394   - Users are strongly encouraged to write "0 < n" rather than
  9395     "n ~= 0". Theorems and proof tools have been modified towards this
  9396     `standard'.
  9397 
  9398 * HOL/Lists:
  9399   the function "set_of_list" has been renamed "set" (and its theorems too);
  9400   the function "nth" now takes its arguments in the reverse order and
  9401   has acquired the infix notation "!" as in "xs!n".
  9402 
  9403 * HOL/Set: UNIV is now a constant and is no longer translated to Compl{};
  9404 
  9405 * HOL/Set: The operator (UN x.B x) now abbreviates (UN x:UNIV. B x) and its
  9406   specialist theorems (like UN1_I) are gone.  Similarly for (INT x.B x);
  9407 
  9408 * HOL/record: extensible records with schematic structural subtyping
  9409 (single inheritance); EXPERIMENTAL version demonstrating the encoding,
  9410 still lacks various theorems and concrete record syntax;
  9411 
  9412 
  9413 *** HOLCF ***
  9414 
  9415 * removed "axioms" and "generated by" sections;
  9416 
  9417 * replaced "ops" section by extended "consts" section, which is capable of
  9418   handling the continuous function space "->" directly;
  9419 
  9420 * domain package:
  9421   . proves theorems immediately and stores them in the theory,
  9422   . creates hierachical name space,
  9423   . now uses normal mixfix annotations (instead of cinfix...),
  9424   . minor changes to some names and values (for consistency),
  9425   . e.g. cases -> casedist, dists_eq -> dist_eqs, [take_lemma] -> take_lemmas,
  9426   . separator between mutual domain defs: changed "," to "and",
  9427   . improved handling of sort constraints;  now they have to
  9428     appear on the left-hand side of the equations only;
  9429 
  9430 * fixed LAM <x,y,zs>.b syntax;
  9431 
  9432 * added extended adm_tac to simplifier in HOLCF -- can now discharge
  9433 adm (%x. P (t x)), where P is chainfinite and t continuous;
  9434 
  9435 
  9436 *** FOL and ZF ***
  9437 
  9438 * FOL: there is a new splitter `split_asm_tac' that can be used e.g.
  9439   with `addloop' of the simplifier to faciliate case splitting in premises.
  9440 
  9441 * qed_spec_mp, qed_goal_spec_mp, qed_goalw_spec_mp are available, as
  9442 in HOL, they strip ALL and --> from proved theorems;
  9443 
  9444 
  9445 
  9446 New in Isabelle94-8 (May 1997)
  9447 ------------------------------
  9448 
  9449 *** General Changes ***
  9450 
  9451 * new utilities to build / run / maintain Isabelle etc. (in parts
  9452 still somewhat experimental); old Makefiles etc. still functional;
  9453 
  9454 * new 'Isabelle System Manual';
  9455 
  9456 * INSTALL text, together with ./configure and ./build scripts;
  9457 
  9458 * reimplemented type inference for greater efficiency, better error
  9459 messages and clean internal interface;
  9460 
  9461 * prlim command for dealing with lots of subgoals (an easier way of
  9462 setting goals_limit);
  9463 
  9464 
  9465 *** Syntax ***
  9466 
  9467 * supports alternative (named) syntax tables (parser and pretty
  9468 printer); internal interface is provided by add_modesyntax(_i);
  9469 
  9470 * Pure, FOL, ZF, HOL, HOLCF now support symbolic input and output; to
  9471 be used in conjunction with the Isabelle symbol font; uses the
  9472 "symbols" syntax table;
  9473 
  9474 * added token_translation interface (may translate name tokens in
  9475 arbitrary ways, dependent on their type (free, bound, tfree, ...) and
  9476 the current print_mode); IMPORTANT: user print translation functions
  9477 are responsible for marking newly introduced bounds
  9478 (Syntax.mark_boundT);
  9479 
  9480 * token translations for modes "xterm" and "xterm_color" that display
  9481 names in bold, underline etc. or colors (which requires a color
  9482 version of xterm);
  9483 
  9484 * infixes may now be declared with names independent of their syntax;
  9485 
  9486 * added typed_print_translation (like print_translation, but may
  9487 access type of constant);
  9488 
  9489 
  9490 *** Classical Reasoner ***
  9491 
  9492 Blast_tac: a new tactic!  It is often more powerful than fast_tac, but has
  9493 some limitations.  Blast_tac...
  9494   + ignores addss, addbefore, addafter; this restriction is intrinsic
  9495   + ignores elimination rules that don't have the correct format
  9496         (the conclusion MUST be a formula variable)
  9497   + ignores types, which can make HOL proofs fail
  9498   + rules must not require higher-order unification, e.g. apply_type in ZF
  9499     [message "Function Var's argument not a bound variable" relates to this]
  9500   + its proof strategy is more general but can actually be slower
  9501 
  9502 * substitution with equality assumptions no longer permutes other
  9503 assumptions;
  9504 
  9505 * minor changes in semantics of addafter (now called addaltern); renamed
  9506 setwrapper to setWrapper and compwrapper to compWrapper; added safe wrapper
  9507 (and access functions for it);
  9508 
  9509 * improved combination of classical reasoner and simplifier:
  9510   + functions for handling clasimpsets
  9511   + improvement of addss: now the simplifier is called _after_ the
  9512     safe steps.
  9513   + safe variant of addss called addSss: uses safe simplifications
  9514     _during_ the safe steps. It is more complete as it allows multiple
  9515     instantiations of unknowns (e.g. with slow_tac).
  9516 
  9517 *** Simplifier ***
  9518 
  9519 * added interface for simplification procedures (functions that
  9520 produce *proven* rewrite rules on the fly, depending on current
  9521 redex);
  9522 
  9523 * ordering on terms as parameter (used for ordered rewriting);
  9524 
  9525 * new functions delcongs, deleqcongs, and Delcongs. richer rep_ss;
  9526 
  9527 * the solver is now split into a safe and an unsafe part.
  9528 This should be invisible for the normal user, except that the
  9529 functions setsolver and addsolver have been renamed to setSolver and
  9530 addSolver; added safe_asm_full_simp_tac;
  9531 
  9532 
  9533 *** HOL ***
  9534 
  9535 * a generic induction tactic `induct_tac' which works for all datatypes and
  9536 also for type `nat';
  9537 
  9538 * a generic case distinction tactic `exhaust_tac' which works for all
  9539 datatypes and also for type `nat';
  9540 
  9541 * each datatype comes with a function `size';
  9542 
  9543 * patterns in case expressions allow tuple patterns as arguments to
  9544 constructors, for example `case x of [] => ... | (x,y,z)#ps => ...';
  9545 
  9546 * primrec now also works with type nat;
  9547 
  9548 * recdef: a new declaration form, allows general recursive functions to be
  9549 defined in theory files.  See HOL/ex/Fib, HOL/ex/Primes, HOL/Subst/Unify.
  9550 
  9551 * the constant for negation has been renamed from "not" to "Not" to
  9552 harmonize with FOL, ZF, LK, etc.;
  9553 
  9554 * HOL/ex/LFilter theory of a corecursive "filter" functional for
  9555 infinite lists;
  9556 
  9557 * HOL/Modelcheck demonstrates invocation of model checker oracle;
  9558 
  9559 * HOL/ex/Ring.thy declares cring_simp, which solves equational
  9560 problems in commutative rings, using axiomatic type classes for + and *;
  9561 
  9562 * more examples in HOL/MiniML and HOL/Auth;
  9563 
  9564 * more default rewrite rules for quantifiers, union/intersection;
  9565 
  9566 * a new constant `arbitrary == @x.False';
  9567 
  9568 * HOLCF/IOA replaces old HOL/IOA;
  9569 
  9570 * HOLCF changes: derived all rules and arities
  9571   + axiomatic type classes instead of classes
  9572   + typedef instead of faking type definitions
  9573   + eliminated the internal constants less_fun, less_cfun, UU_fun, UU_cfun etc.
  9574   + new axclasses cpo, chfin, flat with flat < chfin < pcpo < cpo < po
  9575   + eliminated the types void, one, tr
  9576   + use unit lift and bool lift (with translations) instead of one and tr
  9577   + eliminated blift from Lift3.thy (use Def instead of blift)
  9578   all eliminated rules are derived as theorems --> no visible changes ;
  9579 
  9580 
  9581 *** ZF ***
  9582 
  9583 * ZF now has Fast_tac, Simp_tac and Auto_tac.  Union_iff is a now a default
  9584 rewrite rule; this may affect some proofs.  eq_cs is gone but can be put back
  9585 as ZF_cs addSIs [equalityI];
  9586 
  9587 
  9588 
  9589 New in Isabelle94-7 (November 96)
  9590 ---------------------------------
  9591 
  9592 * allowing negative levels (as offsets) in prlev and choplev;
  9593 
  9594 * super-linear speedup for large simplifications;
  9595 
  9596 * FOL, ZF and HOL now use miniscoping: rewriting pushes
  9597 quantifications in as far as possible (COULD MAKE EXISTING PROOFS
  9598 FAIL); can suppress it using the command Delsimps (ex_simps @
  9599 all_simps); De Morgan laws are also now included, by default;
  9600 
  9601 * improved printing of ==>  :  ~:
  9602 
  9603 * new object-logic "Sequents" adds linear logic, while replacing LK
  9604 and Modal (thanks to Sara Kalvala);
  9605 
  9606 * HOL/Auth: correctness proofs for authentication protocols;
  9607 
  9608 * HOL: new auto_tac combines rewriting and classical reasoning (many
  9609 examples on HOL/Auth);
  9610 
  9611 * HOL: new command AddIffs for declaring theorems of the form P=Q to
  9612 the rewriter and classical reasoner simultaneously;
  9613 
  9614 * function uresult no longer returns theorems in "standard" format;
  9615 regain previous version by: val uresult = standard o uresult;
  9616 
  9617 
  9618 
  9619 New in Isabelle94-6
  9620 -------------------
  9621 
  9622 * oracles -- these establish an interface between Isabelle and trusted
  9623 external reasoners, which may deliver results as theorems;
  9624 
  9625 * proof objects (in particular record all uses of oracles);
  9626 
  9627 * Simp_tac, Fast_tac, etc. that refer to implicit simpset / claset;
  9628 
  9629 * "constdefs" section in theory files;
  9630 
  9631 * "primrec" section (HOL) no longer requires names;
  9632 
  9633 * internal type "tactic" now simply "thm -> thm Sequence.seq";
  9634 
  9635 
  9636 
  9637 New in Isabelle94-5
  9638 -------------------
  9639 
  9640 * reduced space requirements;
  9641 
  9642 * automatic HTML generation from theories;
  9643 
  9644 * theory files no longer require "..." (quotes) around most types;
  9645 
  9646 * new examples, including two proofs of the Church-Rosser theorem;
  9647 
  9648 * non-curried (1994) version of HOL is no longer distributed;
  9649 
  9650 
  9651 
  9652 New in Isabelle94-4
  9653 -------------------
  9654 
  9655 * greatly reduced space requirements;
  9656 
  9657 * theory files (.thy) no longer require \...\ escapes at line breaks;
  9658 
  9659 * searchable theorem database (see the section "Retrieving theorems" on
  9660 page 8 of the Reference Manual);
  9661 
  9662 * new examples, including Grabczewski's monumental case study of the
  9663 Axiom of Choice;
  9664 
  9665 * The previous version of HOL renamed to Old_HOL;
  9666 
  9667 * The new version of HOL (previously called CHOL) uses a curried syntax
  9668 for functions.  Application looks like f a b instead of f(a,b);
  9669 
  9670 * Mutually recursive inductive definitions finally work in HOL;
  9671 
  9672 * In ZF, pattern-matching on tuples is now available in all abstractions and
  9673 translates to the operator "split";
  9674 
  9675 
  9676 
  9677 New in Isabelle94-3
  9678 -------------------
  9679 
  9680 * new infix operator, addss, allowing the classical reasoner to
  9681 perform simplification at each step of its search.  Example:
  9682         fast_tac (cs addss ss)
  9683 
  9684 * a new logic, CHOL, the same as HOL, but with a curried syntax
  9685 for functions.  Application looks like f a b instead of f(a,b).  Also pairs
  9686 look like (a,b) instead of <a,b>;
  9687 
  9688 * PLEASE NOTE: CHOL will eventually replace HOL!
  9689 
  9690 * In CHOL, pattern-matching on tuples is now available in all abstractions.
  9691 It translates to the operator "split".  A new theory of integers is available;
  9692 
  9693 * In ZF, integer numerals now denote two's-complement binary integers.
  9694 Arithmetic operations can be performed by rewriting.  See ZF/ex/Bin.ML;
  9695 
  9696 * Many new examples: I/O automata, Church-Rosser theorem, equivalents
  9697 of the Axiom of Choice;
  9698 
  9699 
  9700 
  9701 New in Isabelle94-2
  9702 -------------------
  9703 
  9704 * Significantly faster resolution;
  9705 
  9706 * the different sections in a .thy file can now be mixed and repeated
  9707 freely;
  9708 
  9709 * Database of theorems for FOL, HOL and ZF.  New
  9710 commands including qed, qed_goal and bind_thm store theorems in the database.
  9711 
  9712 * Simple database queries: return a named theorem (get_thm) or all theorems of
  9713 a given theory (thms_of), or find out what theory a theorem was proved in
  9714 (theory_of_thm);
  9715 
  9716 * Bugs fixed in the inductive definition and datatype packages;
  9717 
  9718 * The classical reasoner provides deepen_tac and depth_tac, making FOL_dup_cs
  9719 and HOL_dup_cs obsolete;
  9720 
  9721 * Syntactic ambiguities caused by the new treatment of syntax in Isabelle94-1
  9722 have been removed;
  9723 
  9724 * Simpler definition of function space in ZF;
  9725 
  9726 * new results about cardinal and ordinal arithmetic in ZF;
  9727 
  9728 * 'subtype' facility in HOL for introducing new types as subsets of existing
  9729 types;
  9730 
  9731 :mode=text:wrap=hard:maxLineLen=72: