src/Pure/more_thm.ML
changeset 49062 7e31dfd99ce7
parent 49058 2924a83a4a0b
child 51316 dfe469293eb4
equal deleted inserted replaced
49061:7449b804073b 49062:7e31dfd99ce7
    93   val lemmaK: string
    93   val lemmaK: string
    94   val corollaryK: string
    94   val corollaryK: string
    95   val legacy_get_kind: thm -> string
    95   val legacy_get_kind: thm -> string
    96   val kind_rule: string -> thm -> thm
    96   val kind_rule: string -> thm -> thm
    97   val kind: string -> attribute
    97   val kind: string -> attribute
    98   val register_proofs: thm list -> Context.generic -> Context.generic
    98   val register_proofs: thm list -> theory -> theory
    99   val register_proofs_thy: thm list -> theory -> theory
       
   100   val begin_proofs: Context.generic -> Context.generic
       
   101   val join_local_proofs: Proof.context -> unit
       
   102   val join_recent_proofs: theory -> unit
       
   103   val join_theory_proofs: theory -> unit
    99   val join_theory_proofs: theory -> unit
   104 end;
   100 end;
   105 
   101 
   106 structure Thm: THM =
   102 structure Thm: THM =
   107 struct
   103 struct
   472 fun kind k = rule_attribute (K (k <> "" ? kind_rule k));
   468 fun kind k = rule_attribute (K (k <> "" ? kind_rule k));
   473 
   469 
   474 
   470 
   475 (* forked proofs *)
   471 (* forked proofs *)
   476 
   472 
   477 type proofs = thm list * unit lazy;  (*accumulated thms, persistent join*)
   473 structure Proofs = Theory_Data
   478 
       
   479 val empty_proofs: proofs = ([], Lazy.value ());
       
   480 
       
   481 fun add_proofs more_thms ((thms, _): proofs) =
       
   482   let val thms' = fold cons more_thms thms
       
   483   in (thms', Lazy.lazy (fn () => Thm.join_proofs (rev thms'))) end;
       
   484 
       
   485 fun force_proofs ((_, prfs): proofs) = Lazy.force prfs;
       
   486 
       
   487 structure Proofs = Generic_Data
       
   488 (
   474 (
   489   type T = proofs * proofs;  (*recent proofs since last begin, all proofs of theory node*)
   475   type T = thm list;
   490   val empty = (empty_proofs, empty_proofs);
   476   val empty = [];
   491   fun extend _ = empty;
   477   fun extend _ = empty;
   492   fun merge _ = empty;
   478   fun merge _ = empty;
   493 );
   479 );
   494 
   480 
   495 val begin_proofs = Proofs.map (apfst (K empty_proofs));
   481 fun register_proofs more_thms = Proofs.map (fn thms => fold cons more_thms thms);
   496 val register_proofs = Proofs.map o pairself o add_proofs;
   482 val join_theory_proofs = Thm.join_proofs o rev o Proofs.get;
   497 val register_proofs_thy = Context.theory_map o register_proofs;
       
   498 
       
   499 val join_local_proofs = force_proofs o #1 o Proofs.get o Context.Proof;
       
   500 val join_recent_proofs = force_proofs o #1 o Proofs.get o Context.Theory;
       
   501 val join_theory_proofs = force_proofs o #2 o Proofs.get o Context.Theory;
       
   502 
   483 
   503 
   484 
   504 open Thm;
   485 open Thm;
   505 
   486 
   506 end;
   487 end;