don't generate 'size_gen_o_map' property if its type variable is too limited anyway to be useful
authorblanchet
Tue, 29 Nov 2016 08:32:46 +0100
changeset 64532 fc2835a932d9
parent 64531 166a2563e0ca
child 64534 ff59fe6b6f6a
don't generate 'size_gen_o_map' property if its type variable is too limited anyway to be useful
NEWS
src/HOL/Tools/BNF/bnf_lfp_size.ML
--- a/NEWS	Tue Nov 29 08:32:44 2016 +0100
+++ b/NEWS	Tue Nov 29 08:32:46 2016 +0100
@@ -6,6 +6,11 @@
 New in this Isabelle version
 ----------------------------
 
+* (Co)datatype package:
+  - The 'size_gen_o_map' lemma is no longer generated for datatypes
+    with type class annotations. As a result, the tactic that derives
+    it no longer fails on nested datatypes. Slight INCOMPATIBILITY.
+
 
 
 New in Isabelle2016-1 (December 2016)
--- a/src/HOL/Tools/BNF/bnf_lfp_size.ML	Tue Nov 29 08:32:44 2016 +0100
+++ b/src/HOL/Tools/BNF/bnf_lfp_size.ML	Tue Nov 29 08:32:46 2016 +0100
@@ -341,7 +341,8 @@
               fold_rev (curry (op @) o #co_rec_o_maps o the o #fp_co_induct_sugar) fp_sugars [];
 
             val size_gen_o_map_thmss =
-              if nested_size_gen_o_maps_complete then
+              if nested_size_gen_o_maps_complete
+                 andalso forall (fn TFree (_, S) => S = @{sort type}) As then
                 @{map 3} (fn goal => fn size_def => fn rec_o_map =>
                     Goal.prove_sorry lthy2 [] [] goal (fn {context = ctxt, ...} =>
                       mk_size_gen_o_map_tac ctxt size_def rec_o_map all_inj_maps nested_size_maps)