| author | krauss | 
| Tue, 26 Oct 2010 13:19:31 +0200 | |
| changeset 40180 | c3ef007115a0 | 
| parent 39557 | fe5722fce758 | 
| child 40627 | becf5d5187cc | 
| permissions | -rw-r--r-- | 
| 31775 | 1 | (* Title: HOL/Tools/Function/size.ML | 
| 29495 | 2 | Author: Stefan Berghofer, Florian Haftmann & Alexander Krauss, TU Muenchen | 
| 24710 | 3 | |
| 4 | Size functions for datatypes. | |
| 5 | *) | |
| 6 | ||
| 7 | signature SIZE = | |
| 8 | sig | |
| 9 | val size_thms: theory -> string -> thm list | |
| 10 | val setup: theory -> theory | |
| 11 | end; | |
| 12 | ||
| 13 | structure Size: SIZE = | |
| 14 | struct | |
| 15 | ||
| 33968 
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
 haftmann parents: 
33766diff
changeset | 16 | open Datatype_Aux; | 
| 24710 | 17 | |
| 33522 | 18 | structure SizeData = Theory_Data | 
| 24714 
1618c2ac1b74
proper Sign operations instead of Theory aliases;
 wenzelm parents: 
24711diff
changeset | 19 | ( | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 20 | type T = (string * thm list) Symtab.table; | 
| 24710 | 21 | val empty = Symtab.empty; | 
| 22 | val extend = I | |
| 33522 | 23 | fun merge data = Symtab.merge (K true) data; | 
| 24714 
1618c2ac1b74
proper Sign operations instead of Theory aliases;
 wenzelm parents: 
24711diff
changeset | 24 | ); | 
| 24710 | 25 | |
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 26 | val lookup_size = SizeData.get #> Symtab.lookup; | 
| 24710 | 27 | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35064diff
changeset | 28 | fun plus (t1, t2) = Const (@{const_name Groups.plus},
 | 
| 24710 | 29 | HOLogic.natT --> HOLogic.natT --> HOLogic.natT) $ t1 $ t2; | 
| 30 | ||
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 31 | fun size_of_type f g h (T as Type (s, Ts)) = | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 32 | (case f s of | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 33 | SOME t => SOME t | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 34 | | NONE => (case g s of | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 35 | SOME size_name => | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 36 | SOME (list_comb (Const (size_name, | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 37 | map (fn U => U --> HOLogic.natT) Ts @ [T] ---> HOLogic.natT), | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 38 | map (size_of_type' f g h) Ts)) | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 39 | | NONE => NONE)) | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 40 | | size_of_type f g h (TFree (s, _)) = h s | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 41 | and size_of_type' f g h T = (case size_of_type f g h T of | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 42 |       NONE => Abs ("x", T, HOLogic.zero)
 | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 43 | | SOME t => t); | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 44 | |
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 45 | fun is_poly thy (DtType (name, dts)) = | 
| 31784 | 46 | (case Datatype.get_info thy name of | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 47 | NONE => false | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 48 | | SOME _ => exists (is_poly thy) dts) | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 49 | | is_poly _ _ = true; | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 50 | |
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 51 | fun constrs_of thy name = | 
| 24710 | 52 | let | 
| 31784 | 53 |     val {descr, index, ...} = Datatype.the_info thy name
 | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 54 | val SOME (_, _, constrs) = AList.lookup op = descr index | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 55 | in constrs end; | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 56 | |
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 57 | val app = curry (list_comb o swap); | 
| 24710 | 58 | |
| 31737 | 59 | fun prove_size_thms (info : info) new_type_names thy = | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 60 | let | 
| 32727 | 61 |     val {descr, alt_names, sorts, rec_names, rec_rewrites, induct, ...} = info;
 | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 62 | val l = length new_type_names; | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 63 | val alt_names' = (case alt_names of | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 64 | NONE => replicate l NONE | SOME names => map SOME names); | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 65 | val descr' = List.take (descr, l); | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 66 | val (rec_names1, rec_names2) = chop l rec_names; | 
| 25864 | 67 | val recTs = get_rec_types descr sorts; | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 68 | val (recTs1, recTs2) = chop l recTs; | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 69 | val (_, (_, paramdts, _)) :: _ = descr; | 
| 25864 | 70 | val paramTs = map (typ_of_dtyp descr sorts) paramdts; | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 71 | val ((param_size_fs, param_size_fTs), f_names) = paramTs |> | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 72 | map (fn T as TFree (s, _) => | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 73 | let | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 74 | val name = "f" ^ implode (tl (explode s)); | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 75 | val U = T --> HOLogic.natT | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 76 | in | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 77 | (((s, Free (name, U)), U), name) | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 78 | end) |> split_list |>> split_list; | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 79 | val param_size = AList.lookup op = param_size_fs; | 
| 24710 | 80 | |
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 81 | val extra_rewrites = descr |> map (#1 o snd) |> distinct op = |> | 
| 29495 | 82 | map_filter (Option.map snd o lookup_size thy) |> flat; | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 83 | val extra_size = Option.map fst o lookup_size thy; | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 84 | |
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 85 | val (((size_names, size_fns), def_names), def_names') = | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 86 | recTs1 ~~ alt_names' |> | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 87 | map (fn (T as Type (s, _), optname) => | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 88 | let | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 89 | val s' = the_default (Long_Name.base_name s) optname ^ "_size"; | 
| 28965 | 90 | val s'' = Sign.full_bname thy s' | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 91 | in | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 92 | (s'', | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 93 | (list_comb (Const (s'', param_size_fTs @ [T] ---> HOLogic.natT), | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 94 | map snd param_size_fs), | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 95 | (s' ^ "_def", s' ^ "_overloaded_def"))) | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 96 | end) |> split_list ||>> split_list ||>> split_list; | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 97 | val overloaded_size_fns = map HOLogic.size_const recTs1; | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 98 | |
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 99 | (* instantiation for primrec combinator *) | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 100 | fun size_of_constr b size_ofp ((_, cargs), (_, cargs')) = | 
| 24710 | 101 | let | 
| 25864 | 102 | val Ts = map (typ_of_dtyp descr sorts) cargs; | 
| 24710 | 103 | val k = length (filter is_rec_type cargs); | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 104 | val (ts, _, _) = fold_rev (fn ((dt, dt'), T) => fn (us, i, j) => | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 105 | if is_rec_type dt then (Bound i :: us, i + 1, j + 1) | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 106 | else | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 107 | (if b andalso is_poly thy dt' then | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 108 | case size_of_type (K NONE) extra_size size_ofp T of | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 109 | NONE => us | SOME sz => sz $ Bound j :: us | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 110 | else us, i, j + 1)) | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 111 | (cargs ~~ cargs' ~~ Ts) ([], 0, k); | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 112 | val t = | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 113 | if null ts andalso (not b orelse not (exists (is_poly thy) cargs')) | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 114 | then HOLogic.zero | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 115 | else foldl1 plus (ts @ [HOLogic.Suc_zero]) | 
| 24710 | 116 | in | 
| 33339 | 117 |         fold_rev (fn T => fn t' => Abs ("x", T, t')) (Ts @ replicate k HOLogic.natT) t
 | 
| 24710 | 118 | end; | 
| 119 | ||
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 120 | val fs = maps (fn (_, (name, _, constrs)) => | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 121 | map (size_of_constr true param_size) (constrs ~~ constrs_of thy name)) descr; | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 122 | val fs' = maps (fn (n, (name, _, constrs)) => | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 123 | map (size_of_constr (l <= n) (K NONE)) (constrs ~~ constrs_of thy name)) descr; | 
| 24710 | 124 | val fTs = map fastype_of fs; | 
| 125 | ||
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 126 | val (rec_combs1, rec_combs2) = chop l (map (fn (T, rec_name) => | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 127 | Const (rec_name, fTs @ [T] ---> HOLogic.natT)) | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 128 | (recTs ~~ rec_names)); | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 129 | |
| 25835 | 130 | fun define_overloaded (def_name, eq) lthy = | 
| 131 | let | |
| 132 | val (Free (c, _), rhs) = (Logic.dest_equals o Syntax.check_term lthy) eq; | |
| 35756 
cfde251d03a5
refining and adding Spec_Rules to definitional packages old_primrec, primrec, recdef, size and function
 bulwahn parents: 
35267diff
changeset | 133 | val (thm, lthy') = lthy | 
| 
cfde251d03a5
refining and adding Spec_Rules to definitional packages old_primrec, primrec, recdef, size and function
 bulwahn parents: 
35267diff
changeset | 134 | |> Local_Theory.define ((Binding.name c, NoSyn), ((Binding.name def_name, []), rhs)) | 
| 
cfde251d03a5
refining and adding Spec_Rules to definitional packages old_primrec, primrec, recdef, size and function
 bulwahn parents: 
35267diff
changeset | 135 | |-> (fn (t, (_, thm)) => Spec_Rules.add Spec_Rules.Equational ([t], [thm]) #> pair thm); | 
| 36610 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 wenzelm parents: 
35756diff
changeset | 136 | val ctxt_thy = ProofContext.init_global (ProofContext.theory_of lthy'); | 
| 25864 | 137 | val thm' = singleton (ProofContext.export lthy' ctxt_thy) thm; | 
| 25835 | 138 | in (thm', lthy') end; | 
| 139 | ||
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 140 | val ((size_def_thms, size_def_thms'), thy') = | 
| 24710 | 141 | thy | 
| 24714 
1618c2ac1b74
proper Sign operations instead of Theory aliases;
 wenzelm parents: 
24711diff
changeset | 142 | |> Sign.add_consts_i (map (fn (s, T) => | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 143 | (Binding.name (Long_Name.base_name s), param_size_fTs @ [T] ---> HOLogic.natT, NoSyn)) | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 144 | (size_names ~~ recTs1)) | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38348diff
changeset | 145 | |> Global_Theory.add_defs false | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 146 | (map (Thm.no_attributes o apsnd (Logic.mk_equals o apsnd (app fs))) | 
| 29579 | 147 | (map Binding.name def_names ~~ (size_fns ~~ rec_combs1))) | 
| 38348 
cf7b2121ad9d
moved instantiation target formally to class_target.ML
 haftmann parents: 
36610diff
changeset | 148 | ||> Class.instantiation | 
| 25890 
0ba401ddbaed
Now uses more carefully designed simpsets to prevent proofs from
 berghofe parents: 
25864diff
changeset | 149 | (map (#1 o snd) descr', map dest_TFree paramTs, [HOLogic.class_size]) | 
| 25835 | 150 | ||>> fold_map define_overloaded | 
| 151 | (def_names' ~~ map Logic.mk_equals (overloaded_size_fns ~~ map (app fs') rec_combs1)) | |
| 152 | ||> Class.prove_instantiation_instance (K (Class.intro_classes_tac [])) | |
| 33671 | 153 | ||> Local_Theory.exit_global; | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 154 | |
| 36610 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 wenzelm parents: 
35756diff
changeset | 155 | val ctxt = ProofContext.init_global thy'; | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 156 | |
| 35064 
1bdef0c013d3
hide fact names clashing with fact names from Group.thy
 haftmann parents: 
35047diff
changeset | 157 |     val simpset1 = HOL_basic_ss addsimps @{thm Nat.add_0} :: @{thm Nat.add_0_right} ::
 | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 158 | size_def_thms @ size_def_thms' @ rec_rewrites @ extra_rewrites; | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 159 | val xs = map (fn i => "x" ^ string_of_int i) (1 upto length recTs2); | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 160 | |
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 161 | fun mk_unfolded_size_eq tab size_ofp fs (p as (x, T), r) = | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 162 | HOLogic.mk_eq (app fs r $ Free p, | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 163 | the (size_of_type tab extra_size size_ofp T) $ Free p); | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 164 | |
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 165 | fun prove_unfolded_size_eqs size_ofp fs = | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 166 | if null recTs2 then [] | 
| 32970 
fbd2bb2489a8
operations of structure Skip_Proof (formerly SkipProof) no longer require quick_and_dirty mode;
 wenzelm parents: 
32957diff
changeset | 167 | else split_conj_thm (Skip_Proof.prove ctxt xs [] | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 168 | (HOLogic.mk_Trueprop (mk_conj (replicate l HOLogic.true_const @ | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 169 | map (mk_unfolded_size_eq (AList.lookup op = | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 170 | (new_type_names ~~ map (app fs) rec_combs1)) size_ofp fs) | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 171 | (xs ~~ recTs2 ~~ rec_combs2)))) | 
| 32712 
ec5976f4d3d8
registering split rules and projected induction rules; ML identifiers more close to Isar theorem names
 haftmann parents: 
31902diff
changeset | 172 | (fn _ => (indtac induct xs THEN_ALL_NEW asm_simp_tac simpset1) 1)); | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 173 | |
| 25890 
0ba401ddbaed
Now uses more carefully designed simpsets to prevent proofs from
 berghofe parents: 
25864diff
changeset | 174 | val unfolded_size_eqs1 = prove_unfolded_size_eqs param_size fs; | 
| 
0ba401ddbaed
Now uses more carefully designed simpsets to prevent proofs from
 berghofe parents: 
25864diff
changeset | 175 | val unfolded_size_eqs2 = prove_unfolded_size_eqs (K NONE) fs'; | 
| 24710 | 176 | |
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 177 | (* characteristic equations for size functions *) | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 178 | fun gen_mk_size_eq p size_of size_ofp size_const T (cname, cargs) = | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 179 | let | 
| 25864 | 180 | val Ts = map (typ_of_dtyp descr sorts) cargs; | 
| 33968 
f94fb13ecbb3
modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
 haftmann parents: 
33766diff
changeset | 181 | val tnames = Name.variant_list f_names (Datatype_Prop.make_tnames Ts); | 
| 29495 | 182 | val ts = map_filter (fn (sT as (s, T), dt) => | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 183 | Option.map (fn sz => sz $ Free sT) | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 184 | (if p dt then size_of_type size_of extra_size size_ofp T | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 185 | else NONE)) (tnames ~~ Ts ~~ cargs) | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 186 | in | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 187 | HOLogic.mk_Trueprop (HOLogic.mk_eq | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 188 | (size_const $ list_comb (Const (cname, Ts ---> T), | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 189 | map2 (curry Free) tnames Ts), | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 190 | if null ts then HOLogic.zero | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 191 | else foldl1 plus (ts @ [HOLogic.Suc_zero]))) | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 192 | end; | 
| 24710 | 193 | |
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 194 | val simpset2 = HOL_basic_ss addsimps | 
| 25890 
0ba401ddbaed
Now uses more carefully designed simpsets to prevent proofs from
 berghofe parents: 
25864diff
changeset | 195 | rec_rewrites @ size_def_thms @ unfolded_size_eqs1; | 
| 
0ba401ddbaed
Now uses more carefully designed simpsets to prevent proofs from
 berghofe parents: 
25864diff
changeset | 196 | val simpset3 = HOL_basic_ss addsimps | 
| 
0ba401ddbaed
Now uses more carefully designed simpsets to prevent proofs from
 berghofe parents: 
25864diff
changeset | 197 | rec_rewrites @ size_def_thms' @ unfolded_size_eqs2; | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 198 | |
| 25890 
0ba401ddbaed
Now uses more carefully designed simpsets to prevent proofs from
 berghofe parents: 
25864diff
changeset | 199 | fun prove_size_eqs p size_fns size_ofp simpset = | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 200 | maps (fn (((_, (_, _, constrs)), size_const), T) => | 
| 35021 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
34974diff
changeset | 201 | map (fn constr => Drule.export_without_context (Skip_Proof.prove ctxt [] [] | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 202 | (gen_mk_size_eq p (AList.lookup op = (new_type_names ~~ size_fns)) | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 203 | size_ofp size_const T constr) | 
| 25890 
0ba401ddbaed
Now uses more carefully designed simpsets to prevent proofs from
 berghofe parents: 
25864diff
changeset | 204 | (fn _ => simp_tac simpset 1))) constrs) | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 205 | (descr' ~~ size_fns ~~ recTs1); | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 206 | |
| 25890 
0ba401ddbaed
Now uses more carefully designed simpsets to prevent proofs from
 berghofe parents: 
25864diff
changeset | 207 | val size_eqns = prove_size_eqs (is_poly thy') size_fns param_size simpset2 @ | 
| 
0ba401ddbaed
Now uses more carefully designed simpsets to prevent proofs from
 berghofe parents: 
25864diff
changeset | 208 | prove_size_eqs is_rec_type overloaded_size_fns (K NONE) simpset3; | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 209 | |
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38348diff
changeset | 210 | val ([size_thms], thy'') = | 
| 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38348diff
changeset | 211 | Global_Theory.add_thmss | 
| 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38348diff
changeset | 212 | [((Binding.name "size", size_eqns), | 
| 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38348diff
changeset | 213 | [Simplifier.simp_add, Nitpick_Simps.add, | 
| 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38348diff
changeset | 214 | Thm.declaration_attribute | 
| 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38348diff
changeset | 215 | (fn thm => Context.mapping (Code.add_default_eqn thm) I)])] thy'; | 
| 24710 | 216 | |
| 217 | in | |
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 218 | SizeData.map (fold (Symtab.update_new o apsnd (rpair size_thms)) | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 219 | (new_type_names ~~ size_names)) thy'' | 
| 24710 | 220 | end; | 
| 221 | ||
| 31668 
a616e56a5ec8
datatype packages: record datatype_config for configuration flags; less verbose signatures
 haftmann parents: 
30364diff
changeset | 222 | fun add_size_thms config (new_type_names as name :: _) thy = | 
| 24710 | 223 | let | 
| 31784 | 224 |     val info as {descr, alt_names, ...} = Datatype.the_info thy name;
 | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 225 | val prefix = Long_Name.map_base_name (K (space_implode "_" | 
| 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 226 | (the_default (map Long_Name.base_name new_type_names) alt_names))) name; | 
| 24710 | 227 | val no_size = exists (fn (_, (_, _, constrs)) => exists (fn (_, cargs) => exists (fn dt => | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 228 | is_rec_type dt andalso not (null (fst (strip_dtyp dt)))) cargs) constrs) descr | 
| 24710 | 229 | in if no_size then thy | 
| 230 | else | |
| 231 | thy | |
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 232 | |> Sign.root_path | 
| 24714 
1618c2ac1b74
proper Sign operations instead of Theory aliases;
 wenzelm parents: 
24711diff
changeset | 233 | |> Sign.add_path prefix | 
| 28361 
232fcbba2e4e
explicit checkpoint for low-level (global) theory operations, admits concurrent proofs;
 wenzelm parents: 
28083diff
changeset | 234 | |> Theory.checkpoint | 
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 235 | |> prove_size_thms info new_type_names | 
| 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 236 | |> Sign.restore_naming thy | 
| 24710 | 237 | end; | 
| 238 | ||
| 25679 
b77f797b528a
size functions for nested datatypes are now expressed using
 berghofe parents: 
25569diff
changeset | 239 | val size_thms = snd oo (the oo lookup_size); | 
| 24710 | 240 | |
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31668diff
changeset | 241 | val setup = Datatype.interpretation add_size_thms; | 
| 24710 | 242 | |
| 29866 
6e93ae65c678
Added Nitpick_Const_Psimp attribute, dropped the 's' in Nitpick_Const_Simps, and killed the Nitpick_Ind_Intros attribute.
 blanchet parents: 
29863diff
changeset | 243 | end; |