src/Pure/axclass.ML
changeset 27547 13199740ced6
parent 27497 c683836fe908
child 27554 2deaa546ba0d
equal deleted inserted replaced
27546:726e8fa3e404 27547:13199740ced6
   180     SOME th => Thm.transfer thy th
   180     SOME th => Thm.transfer thy th
   181   | NONE => error ("Unproven type arity " ^
   181   | NONE => error ("Unproven type arity " ^
   182       Syntax.string_of_arity (ProofContext.init thy) (a, Ss, [c])));
   182       Syntax.string_of_arity (ProofContext.init thy) (a, Ss, [c])));
   183 
   183 
   184 fun arity_property thy (c, a) x =
   184 fun arity_property thy (c, a) x =
   185   these (Symtab.lookup (snd (get_instances thy)) a)
   185   Symtab.lookup_list (snd (get_instances thy)) a
   186   |> map_filter (fn ((c', _), th) => if c = c'
   186   |> map_filter (fn ((c', _), th) => if c = c'
   187         then AList.lookup (op =) (Thm.get_tags th) x else NONE)
   187       then AList.lookup (op =) (Thm.get_tags th) x else NONE)
   188   |> rev;
   188   |> rev;
   189 
   189 
   190 fun put_arity ((t, Ss, c), th) = map_instances (fn (classrel, arities) =>
   190 fun insert_arity_completions thy (t, ((c, Ss), th)) arities =
   191   (classrel, arities |> Symtab.insert_list (eq_fst op =) (t, ((c, Ss), th))));
   191   let
   192 
   192     val algebra = Sign.classes_of thy;
   193 
   193     val super_class_completions = Sign.super_classes thy c
       
   194       |> filter_out (fn c1 => exists (fn ((c2, Ss2), _) => c1 = c2
       
   195           andalso Sorts.sorts_le algebra (Ss2, Ss)) (Symtab.lookup_list arities t))
       
   196     val tags = Thm.get_tags th;
       
   197     val completions = map (fn c1 => (Sorts.weaken ((#classes o Sorts.rep_algebra) algebra)
       
   198       (fn (th, c2) => fn c3 => th RS the_classrel thy (c2, c3)) (th, c) c1
       
   199         |> Thm.map_tags (K tags) |> Thm.close_derivation, c1)) super_class_completions;
       
   200     val arities' = fold (fn (th1, c1) => Symtab.cons_list (t, ((c1, Ss), th1)))
       
   201       completions arities;
       
   202   in (completions, arities') end;
       
   203 
       
   204 fun put_arity ((t, Ss, c), th) thy =
       
   205   let
       
   206     val th' = (Thm.map_tags o AList.default (op =))
       
   207       (Markup.theory_nameN, Context.theory_name thy) th;
       
   208     val arity' = (t, ((c, Ss), th'));
       
   209   in
       
   210     thy
       
   211     |> map_instances (fn (classrel, arities) => (classrel,
       
   212       arities
       
   213       |> Symtab.insert_list (eq_fst op =) arity'
       
   214       |> insert_arity_completions thy arity'
       
   215       |> snd))
       
   216   end;
       
   217 
       
   218 fun complete_arities thy = 
       
   219   let
       
   220     val arities = snd (get_instances thy);
       
   221     val (completions, arities') = arities
       
   222       |> fold_map (insert_arity_completions thy) (Symtab.dest_list arities)
       
   223       |>> flat;
       
   224   in if null completions
       
   225     then NONE
       
   226     else SOME (thy |> map_instances (fn (classrel, _) => (classrel, arities')))
       
   227   end;
       
   228 
       
   229 val _ = Context.>> (Context.map_theory (Theory.at_begin complete_arities));
   194 
   230 
   195 
   231 
   196 (* maintain instance parameters *)
   232 (* maintain instance parameters *)
   197 
   233 
   198 val get_inst_params = #2 o #2 o AxClassData.get;
   234 val get_inst_params = #2 o #2 o AxClassData.get;
   270     val _ = map (Sign.certify_sort thy) Ss = Ss orelse err ();
   306     val _ = map (Sign.certify_sort thy) Ss = Ss orelse err ();
   271     val _ = case filter_out (fn c => can (get_inst_param thy) (c, t)) (params_of thy c)
   307     val _ = case filter_out (fn c => can (get_inst_param thy) (c, t)) (params_of thy c)
   272      of [] => ()
   308      of [] => ()
   273       | cs => Output.legacy_feature
   309       | cs => Output.legacy_feature
   274           ("Missing specifications for overloaded parameters " ^ commas_quote cs)
   310           ("Missing specifications for overloaded parameters " ^ commas_quote cs)
   275     val th' = th
   311     val th' = Drule.unconstrainTs th;
   276       |> Drule.unconstrainTs
       
   277       |> (Thm.map_tags o AList.default (op =))
       
   278           (Markup.theory_nameN, Context.theory_name thy)
       
   279   in
   312   in
   280     thy
   313     thy
   281     |> Sign.primitive_arity (t, Ss, [c])
   314     |> Sign.primitive_arity (t, Ss, [c])
   282     |> put_arity ((t, Ss, c), th')
   315     |> put_arity ((t, Ss, c), th')
   283   end;
   316   end;