| author | Fabian Huch <huch@in.tum.de> | 
| Tue, 10 Oct 2023 13:15:32 +0200 | |
| changeset 78747 | d03de7bc2137 | 
| parent 78453 | 3fdf3c5cfa9d | 
| child 80306 | c2537860ccf8 | 
| permissions | -rw-r--r-- | 
| 
18140
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
1  | 
(* Title: Pure/Isar/element.ML  | 
| 
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
Author: Makarius  | 
| 
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
|
| 19777 | 4  | 
Explicit data structures for some Isar language elements, with derived  | 
5  | 
logical operations.  | 
|
| 
18140
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
6  | 
*)  | 
| 
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
signature ELEMENT =  | 
| 
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
sig  | 
| 60448 | 10  | 
  type ('typ, 'term) obtain = binding * ((binding * 'typ option * mixfix) list * 'term list)
 | 
11  | 
type obtains = (string, string) obtain list  | 
|
12  | 
type obtains_i = (typ, term) obtain list  | 
|
| 19259 | 13  | 
  datatype ('typ, 'term) stmt =
 | 
| 
28084
 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 
wenzelm 
parents: 
28079 
diff
changeset
 | 
14  | 
    Shows of (Attrib.binding * ('term * 'term list) list) list |
 | 
| 60444 | 15  | 
    Obtains of ('typ, 'term) obtain list
 | 
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25739 
diff
changeset
 | 
16  | 
type statement = (string, string) stmt  | 
| 
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25739 
diff
changeset
 | 
17  | 
type statement_i = (typ, term) stmt  | 
| 
18140
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
18  | 
  datatype ('typ, 'term, 'fact) ctxt =
 | 
| 29578 | 19  | 
Fixes of (binding * 'typ option * mixfix) list |  | 
| 
18140
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
20  | 
Constrains of (string * 'typ) list |  | 
| 
28084
 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 
wenzelm 
parents: 
28079 
diff
changeset
 | 
21  | 
    Assumes of (Attrib.binding * ('term * 'term list) list) list |
 | 
| 
 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 
wenzelm 
parents: 
28079 
diff
changeset
 | 
22  | 
    Defines of (Attrib.binding * ('term * 'term list)) list |
 | 
| 
67671
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
23  | 
    Notes of string * (Attrib.binding * ('fact * Token.src list) list) list |
 | 
| 
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
24  | 
Lazy_Notes of string * (binding * 'fact lazy)  | 
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25739 
diff
changeset
 | 
25  | 
type context = (string, string, Facts.ref) ctxt  | 
| 
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25739 
diff
changeset
 | 
26  | 
type context_i = (typ, term, thm list) ctxt  | 
| 29603 | 27  | 
  val map_ctxt: {binding: binding -> binding, typ: 'typ -> 'a, term: 'term -> 'b,
 | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
58002 
diff
changeset
 | 
28  | 
pattern: 'term -> 'b, fact: 'fact -> 'c, attrib: Token.src -> Token.src} ->  | 
| 29603 | 29  | 
    ('typ, 'term, 'fact) ctxt -> ('a, 'b, 'c) ctxt
 | 
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
58002 
diff
changeset
 | 
30  | 
val map_ctxt_attrib: (Token.src -> Token.src) ->  | 
| 21528 | 31  | 
    ('typ, 'term, 'fact) ctxt -> ('typ, 'term, 'fact) ctxt
 | 
| 
78064
 
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
 
wenzelm 
parents: 
78062 
diff
changeset
 | 
32  | 
val trim_context_ctxt: context_i -> context_i  | 
| 
 
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
 
wenzelm 
parents: 
78062 
diff
changeset
 | 
33  | 
val transfer_ctxt: theory -> context_i -> context_i  | 
| 45290 | 34  | 
val transform_ctxt: morphism -> context_i -> context_i  | 
| 19777 | 35  | 
val pretty_stmt: Proof.context -> statement_i -> Pretty.T list  | 
36  | 
val pretty_ctxt: Proof.context -> context_i -> Pretty.T list  | 
|
| 59385 | 37  | 
val pretty_ctxt_no_attribs: Proof.context -> context_i -> Pretty.T list  | 
| 19777 | 38  | 
val pretty_statement: Proof.context -> string -> thm -> Pretty.T  | 
39  | 
type witness  | 
|
| 29578 | 40  | 
val prove_witness: Proof.context -> term -> tactic -> witness  | 
41  | 
val witness_proof: (witness list list -> Proof.context -> Proof.context) ->  | 
|
42  | 
term list list -> Proof.context -> Proof.state  | 
|
43  | 
val witness_proof_eqs: (witness list list -> thm list -> Proof.context -> Proof.context) ->  | 
|
44  | 
term list list -> term list -> Proof.context -> Proof.state  | 
|
45  | 
val witness_local_proof: (witness list list -> Proof.state -> Proof.state) ->  | 
|
| 
62680
 
646b84666a56
eliminated unused argument (see also 58110c1e02bc);
 
wenzelm 
parents: 
62094 
diff
changeset
 | 
46  | 
string -> term list list -> Proof.context -> Proof.state -> Proof.state  | 
| 38108 | 47  | 
val witness_local_proof_eqs: (witness list list -> thm list -> Proof.state -> Proof.state) ->  | 
| 
62680
 
646b84666a56
eliminated unused argument (see also 58110c1e02bc);
 
wenzelm 
parents: 
62094 
diff
changeset
 | 
48  | 
string -> term list list -> term list -> Proof.context -> Proof.state -> Proof.state  | 
| 45290 | 49  | 
val transform_witness: morphism -> witness -> witness  | 
| 
54883
 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 
wenzelm 
parents: 
54742 
diff
changeset
 | 
50  | 
val conclude_witness: Proof.context -> witness -> thm  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22568 
diff
changeset
 | 
51  | 
val pretty_witness: Proof.context -> witness -> Pretty.T  | 
| 74282 | 52  | 
val instantiate_normalize_morphism: ctyp TFrees.table * cterm Frees.table -> morphism  | 
| 21481 | 53  | 
val satisfy_morphism: witness list -> morphism  | 
| 
78453
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
54  | 
val eq_term_morphism: Proof.context -> term list -> morphism option  | 
| 
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
55  | 
val eq_morphism: Proof.context -> thm list -> morphism option  | 
| 
30775
 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
56  | 
val init: context_i -> Context.generic -> Context.generic  | 
| 
30777
 
9960ff945c52
simplified Element.activate(_i): singleton version;
 
wenzelm 
parents: 
30775 
diff
changeset
 | 
57  | 
val activate_i: context_i -> Proof.context -> context_i * Proof.context  | 
| 
 
9960ff945c52
simplified Element.activate(_i): singleton version;
 
wenzelm 
parents: 
30775 
diff
changeset
 | 
58  | 
val activate: (typ, term, Facts.ref) ctxt -> Proof.context -> context_i * Proof.context  | 
| 
18140
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
59  | 
end;  | 
| 
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
60  | 
|
| 
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
61  | 
structure Element: ELEMENT =  | 
| 
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
62  | 
struct  | 
| 
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
63  | 
|
| 19777 | 64  | 
(** language elements **)  | 
65  | 
||
66  | 
(* statement *)  | 
|
| 19259 | 67  | 
|
| 60448 | 68  | 
type ('typ, 'term) obtain = binding * ((binding * 'typ option * mixfix) list * 'term list);
 | 
69  | 
type obtains = (string, string) obtain list;  | 
|
70  | 
type obtains_i = (typ, term) obtain list;  | 
|
71  | 
||
| 19259 | 72  | 
datatype ('typ, 'term) stmt =
 | 
| 
28084
 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 
wenzelm 
parents: 
28079 
diff
changeset
 | 
73  | 
  Shows of (Attrib.binding * ('term * 'term list) list) list |
 | 
| 60444 | 74  | 
  Obtains of ('typ, 'term) obtain list;
 | 
| 19259 | 75  | 
|
76  | 
type statement = (string, string) stmt;  | 
|
77  | 
type statement_i = (typ, term) stmt;  | 
|
78  | 
||
79  | 
||
| 19777 | 80  | 
(* context *)  | 
| 
18140
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
81  | 
|
| 
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
82  | 
datatype ('typ, 'term, 'fact) ctxt =
 | 
| 29578 | 83  | 
Fixes of (binding * 'typ option * mixfix) list |  | 
| 
18140
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
84  | 
Constrains of (string * 'typ) list |  | 
| 
28084
 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 
wenzelm 
parents: 
28079 
diff
changeset
 | 
85  | 
  Assumes of (Attrib.binding * ('term * 'term list) list) list |
 | 
| 
 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 
wenzelm 
parents: 
28079 
diff
changeset
 | 
86  | 
  Defines of (Attrib.binding * ('term * 'term list)) list |
 | 
| 
67671
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
87  | 
  Notes of string * (Attrib.binding * ('fact * Token.src list) list) list |
 | 
| 
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
88  | 
Lazy_Notes of string * (binding * 'fact lazy);  | 
| 
18140
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
89  | 
|
| 
26336
 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 
wenzelm 
parents: 
25739 
diff
changeset
 | 
90  | 
type context = (string, string, Facts.ref) ctxt;  | 
| 
18140
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
91  | 
type context_i = (typ, term, thm list) ctxt;  | 
| 
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
92  | 
|
| 29603 | 93  | 
fun map_ctxt {binding, typ, term, pattern, fact, attrib} =
 | 
94  | 
fn Fixes fixes => Fixes (fixes |> map (fn (x, T, mx) => (binding x, Option.map typ T, mx)))  | 
|
| 
28079
 
955c42c8a5e4
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
27865 
diff
changeset
 | 
95  | 
| Constrains xs => Constrains (xs |> map (fn (x, T) =>  | 
| 42494 | 96  | 
(Variable.check_name (binding (Binding.name x)), typ T)))  | 
| 
18140
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
97  | 
| Assumes asms => Assumes (asms |> map (fn ((a, atts), propps) =>  | 
| 29603 | 98  | 
((binding a, map attrib atts), propps |> map (fn (t, ps) => (term t, map pattern ps)))))  | 
| 
18140
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
99  | 
| Defines defs => Defines (defs |> map (fn ((a, atts), (t, ps)) =>  | 
| 29603 | 100  | 
((binding a, map attrib atts), (term t, map pattern ps))))  | 
| 21440 | 101  | 
| Notes (kind, facts) => Notes (kind, facts |> map (fn ((a, atts), bs) =>  | 
| 
67671
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
102  | 
((binding a, map attrib atts), bs |> map (fn (ths, btts) => (fact ths, map attrib btts)))))  | 
| 
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
103  | 
| Lazy_Notes (kind, (a, ths)) => Lazy_Notes (kind, (binding a, Lazy.map fact ths));  | 
| 
18140
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
104  | 
|
| 21528 | 105  | 
fun map_ctxt_attrib attrib =  | 
| 29603 | 106  | 
  map_ctxt {binding = I, typ = I, term = I, pattern = I, fact = I, attrib = attrib};
 | 
| 21528 | 107  | 
|
| 
78064
 
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
 
wenzelm 
parents: 
78062 
diff
changeset
 | 
108  | 
val trim_context_ctxt: context_i -> context_i = map_ctxt  | 
| 
 
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
 
wenzelm 
parents: 
78062 
diff
changeset
 | 
109  | 
 {binding = I, typ = I, term = I, pattern = I,
 | 
| 
 
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
 
wenzelm 
parents: 
78062 
diff
changeset
 | 
110  | 
fact = map Thm.trim_context,  | 
| 
 
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
 
wenzelm 
parents: 
78062 
diff
changeset
 | 
111  | 
attrib = map Token.trim_context};  | 
| 
 
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
 
wenzelm 
parents: 
78062 
diff
changeset
 | 
112  | 
|
| 
 
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
 
wenzelm 
parents: 
78062 
diff
changeset
 | 
113  | 
fun transfer_ctxt thy: context_i -> context_i = map_ctxt  | 
| 
 
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
 
wenzelm 
parents: 
78062 
diff
changeset
 | 
114  | 
 {binding = I, typ = I, term = I, pattern = I,
 | 
| 
 
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
 
wenzelm 
parents: 
78062 
diff
changeset
 | 
115  | 
fact = map (Thm.transfer thy),  | 
| 
 
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
 
wenzelm 
parents: 
78062 
diff
changeset
 | 
116  | 
attrib = map (Token.transfer thy)};  | 
| 
 
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
 
wenzelm 
parents: 
78062 
diff
changeset
 | 
117  | 
|
| 45290 | 118  | 
fun transform_ctxt phi = map_ctxt  | 
| 28965 | 119  | 
 {binding = Morphism.binding phi,
 | 
| 21481 | 120  | 
typ = Morphism.typ phi,  | 
121  | 
term = Morphism.term phi,  | 
|
| 29603 | 122  | 
pattern = Morphism.term phi,  | 
| 21521 | 123  | 
fact = Morphism.fact phi,  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
124  | 
attrib = map (Token.transform phi)};  | 
| 
18140
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
125  | 
|
| 19808 | 126  | 
|
| 18894 | 127  | 
|
| 19259 | 128  | 
(** pretty printing **)  | 
129  | 
||
| 19267 | 130  | 
fun pretty_items _ _ [] = []  | 
131  | 
| pretty_items keyword sep (x :: ys) =  | 
|
| 55763 | 132  | 
Pretty.block [Pretty.keyword2 keyword, Pretty.brk 1, x] ::  | 
133  | 
map (fn y => Pretty.block [Pretty.str " ", Pretty.keyword2 sep, Pretty.brk 1, y]) ys;  | 
|
| 19259 | 134  | 
|
135  | 
||
136  | 
(* pretty_stmt *)  | 
|
137  | 
||
138  | 
fun pretty_stmt ctxt =  | 
|
139  | 
let  | 
|
| 24920 | 140  | 
val prt_typ = Pretty.quote o Syntax.pretty_typ ctxt;  | 
141  | 
val prt_term = Pretty.quote o Syntax.pretty_term ctxt;  | 
|
| 55763 | 142  | 
val prt_terms = separate (Pretty.keyword2 "and") o map prt_term;  | 
| 
60242
 
3a8501876dba
tuned output -- avoid empty quites and extra breaks;
 
wenzelm 
parents: 
59970 
diff
changeset
 | 
143  | 
val prt_binding = Attrib.pretty_binding ctxt;  | 
| 
64398
 
5076725247fa
more robust printing of names in the context of outer syntax;
 
wenzelm 
parents: 
63395 
diff
changeset
 | 
144  | 
val prt_name = Proof_Context.pretty_name ctxt;  | 
| 19259 | 145  | 
|
146  | 
fun prt_show (a, ts) =  | 
|
| 
60242
 
3a8501876dba
tuned output -- avoid empty quites and extra breaks;
 
wenzelm 
parents: 
59970 
diff
changeset
 | 
147  | 
Pretty.block (Pretty.breaks (prt_binding a ":" @ prt_terms (map fst ts)));  | 
| 19259 | 148  | 
|
| 60448 | 149  | 
fun prt_var (x, SOME T, _) = Pretty.block  | 
| 
64398
 
5076725247fa
more robust printing of names in the context of outer syntax;
 
wenzelm 
parents: 
63395 
diff
changeset
 | 
150  | 
[prt_name (Binding.name_of x), Pretty.str " ::", Pretty.brk 1, prt_typ T]  | 
| 
 
5076725247fa
more robust printing of names in the context of outer syntax;
 
wenzelm 
parents: 
63395 
diff
changeset
 | 
151  | 
| prt_var (x, NONE, _) = prt_name (Binding.name_of x);  | 
| 55763 | 152  | 
val prt_vars = separate (Pretty.keyword2 "and") o map prt_var;  | 
| 19259 | 153  | 
|
| 60448 | 154  | 
fun prt_obtain (_, ([], props)) = Pretty.block (Pretty.breaks (prt_terms props))  | 
155  | 
| prt_obtain (_, (vars, props)) = Pretty.block (Pretty.breaks  | 
|
156  | 
(prt_vars vars @ [Pretty.keyword2 "where"] @ prt_terms props));  | 
|
| 19259 | 157  | 
in  | 
| 19267 | 158  | 
fn Shows shows => pretty_items "shows" "and" (map prt_show shows)  | 
159  | 
| Obtains obtains => pretty_items "obtains" "|" (map prt_obtain obtains)  | 
|
| 19259 | 160  | 
end;  | 
161  | 
||
| 18894 | 162  | 
|
| 19259 | 163  | 
(* pretty_ctxt *)  | 
164  | 
||
| 59385 | 165  | 
fun gen_pretty_ctxt show_attribs ctxt =  | 
| 19259 | 166  | 
let  | 
| 24920 | 167  | 
val prt_typ = Pretty.quote o Syntax.pretty_typ ctxt;  | 
168  | 
val prt_term = Pretty.quote o Syntax.pretty_term ctxt;  | 
|
| 62094 | 169  | 
val prt_thm = Pretty.cartouche o Thm.pretty_thm ctxt;  | 
| 
64398
 
5076725247fa
more robust printing of names in the context of outer syntax;
 
wenzelm 
parents: 
63395 
diff
changeset
 | 
170  | 
val prt_name = Proof_Context.pretty_name ctxt;  | 
| 59385 | 171  | 
|
| 
60242
 
3a8501876dba
tuned output -- avoid empty quites and extra breaks;
 
wenzelm 
parents: 
59970 
diff
changeset
 | 
172  | 
fun prt_binding (b, atts) =  | 
| 
 
3a8501876dba
tuned output -- avoid empty quites and extra breaks;
 
wenzelm 
parents: 
59970 
diff
changeset
 | 
173  | 
Attrib.pretty_binding ctxt (b, if show_attribs then atts else []);  | 
| 59385 | 174  | 
|
175  | 
fun prt_fact (ths, atts) =  | 
|
176  | 
if not show_attribs orelse null atts then map prt_thm ths  | 
|
177  | 
else  | 
|
178  | 
        Pretty.enclose "(" ")" (Pretty.breaks (map prt_thm ths)) ::
 | 
|
179  | 
Attrib.pretty_attribs ctxt atts;  | 
|
| 19259 | 180  | 
|
| 19267 | 181  | 
fun prt_mixfix NoSyn = []  | 
| 
42287
 
d98eb048a2e4
discontinued special treatment of structure Mixfix;
 
wenzelm 
parents: 
41581 
diff
changeset
 | 
182  | 
| prt_mixfix mx = [Pretty.brk 2, Mixfix.pretty_mixfix mx];  | 
| 19267 | 183  | 
|
| 
64398
 
5076725247fa
more robust printing of names in the context of outer syntax;
 
wenzelm 
parents: 
63395 
diff
changeset
 | 
184  | 
fun prt_fix (x, SOME T, mx) = Pretty.block (prt_name (Binding.name_of x) :: Pretty.str " ::" ::  | 
| 68274 | 185  | 
Pretty.brk 1 :: prt_typ T :: prt_mixfix mx)  | 
186  | 
| prt_fix (x, NONE, mx) = Pretty.block (prt_name (Binding.name_of x) :: prt_mixfix mx);  | 
|
| 28965 | 187  | 
fun prt_constrain (x, T) = prt_fix (Binding.name x, SOME T, NoSyn);  | 
| 18894 | 188  | 
|
| 19259 | 189  | 
fun prt_asm (a, ts) =  | 
| 
60242
 
3a8501876dba
tuned output -- avoid empty quites and extra breaks;
 
wenzelm 
parents: 
59970 
diff
changeset
 | 
190  | 
Pretty.block (Pretty.breaks (prt_binding a ":" @ map (prt_term o fst) ts));  | 
| 19259 | 191  | 
fun prt_def (a, (t, _)) =  | 
| 
60242
 
3a8501876dba
tuned output -- avoid empty quites and extra breaks;
 
wenzelm 
parents: 
59970 
diff
changeset
 | 
192  | 
Pretty.block (Pretty.breaks (prt_binding a ":" @ [prt_term t]));  | 
| 19259 | 193  | 
|
194  | 
fun prt_note (a, ths) =  | 
|
| 
60242
 
3a8501876dba
tuned output -- avoid empty quites and extra breaks;
 
wenzelm 
parents: 
59970 
diff
changeset
 | 
195  | 
Pretty.block (Pretty.breaks (flat (prt_binding a " =" :: map prt_fact ths)));  | 
| 
67671
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
196  | 
|
| 
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
197  | 
fun notes_kind "" = "notes"  | 
| 
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
198  | 
| notes_kind kind = "notes " ^ kind;  | 
| 19259 | 199  | 
in  | 
| 19267 | 200  | 
fn Fixes fixes => pretty_items "fixes" "and" (map prt_fix fixes)  | 
201  | 
| Constrains xs => pretty_items "constrains" "and" (map prt_constrain xs)  | 
|
202  | 
| Assumes asms => pretty_items "assumes" "and" (map prt_asm asms)  | 
|
203  | 
| Defines defs => pretty_items "defines" "and" (map prt_def defs)  | 
|
| 
67671
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
204  | 
| Notes (kind, facts) => pretty_items (notes_kind kind) "and" (map prt_note facts)  | 
| 
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
205  | 
| Lazy_Notes (kind, (a, ths)) =>  | 
| 
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
206  | 
pretty_items (notes_kind kind) "and" [prt_note ((a, []), [(Lazy.force ths, [])])]  | 
| 19259 | 207  | 
end;  | 
| 18894 | 208  | 
|
| 59385 | 209  | 
val pretty_ctxt = gen_pretty_ctxt true;  | 
210  | 
val pretty_ctxt_no_attribs = gen_pretty_ctxt false;  | 
|
211  | 
||
| 19267 | 212  | 
|
213  | 
(* pretty_statement *)  | 
|
214  | 
||
215  | 
local  | 
|
216  | 
||
| 59970 | 217  | 
fun standard_elim ctxt th =  | 
218  | 
(case Object_Logic.elim_concl ctxt th of  | 
|
| 
41581
 
72a02e3dec7e
clarified pretty_statement: more robust treatment of fixes and conclusion of elimination (e.g. for classical rule);
 
wenzelm 
parents: 
41425 
diff
changeset
 | 
219  | 
SOME C =>  | 
| 
 
72a02e3dec7e
clarified pretty_statement: more robust treatment of fixes and conclusion of elimination (e.g. for classical rule);
 
wenzelm 
parents: 
41425 
diff
changeset
 | 
220  | 
let  | 
| 
 
72a02e3dec7e
clarified pretty_statement: more robust treatment of fixes and conclusion of elimination (e.g. for classical rule);
 
wenzelm 
parents: 
41425 
diff
changeset
 | 
221  | 
val thesis = Var ((Auto_Bind.thesisN, Thm.maxidx_of th + 1), fastype_of C);  | 
| 77879 | 222  | 
val insts = (TVars.empty, Vars.make1 (Term.dest_Var C, Thm.cterm_of ctxt thesis));  | 
| 74282 | 223  | 
val th' = Thm.instantiate insts th;  | 
| 
41581
 
72a02e3dec7e
clarified pretty_statement: more robust treatment of fixes and conclusion of elimination (e.g. for classical rule);
 
wenzelm 
parents: 
41425 
diff
changeset
 | 
224  | 
in (th', true) end  | 
| 
 
72a02e3dec7e
clarified pretty_statement: more robust treatment of fixes and conclusion of elimination (e.g. for classical rule);
 
wenzelm 
parents: 
41425 
diff
changeset
 | 
225  | 
| NONE => (th, false));  | 
| 
 
72a02e3dec7e
clarified pretty_statement: more robust treatment of fixes and conclusion of elimination (e.g. for classical rule);
 
wenzelm 
parents: 
41425 
diff
changeset
 | 
226  | 
|
| 
64398
 
5076725247fa
more robust printing of names in the context of outer syntax;
 
wenzelm 
parents: 
63395 
diff
changeset
 | 
227  | 
fun thm_name ctxt kind th prts =  | 
| 19267 | 228  | 
let val head =  | 
| 
27865
 
27a8ad9612a3
moved basic thm operations from structure PureThy to Thm (cf. more_thm.ML);
 
wenzelm 
parents: 
26721 
diff
changeset
 | 
229  | 
if Thm.has_name_hint th then  | 
| 
64398
 
5076725247fa
more robust printing of names in the context of outer syntax;
 
wenzelm 
parents: 
63395 
diff
changeset
 | 
230  | 
Pretty.block [Pretty.keyword1 kind, Pretty.brk 1,  | 
| 
 
5076725247fa
more robust printing of names in the context of outer syntax;
 
wenzelm 
parents: 
63395 
diff
changeset
 | 
231  | 
Proof_Context.pretty_name ctxt (Long_Name.base_name (Thm.get_name_hint th)), Pretty.str ":"]  | 
| 55763 | 232  | 
else Pretty.keyword1 kind  | 
| 19267 | 233  | 
in Pretty.block (Pretty.fbreaks (head :: prts)) end;  | 
234  | 
||
235  | 
fun obtain prop ctxt =  | 
|
236  | 
let  | 
|
| 
60695
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60642 
diff
changeset
 | 
237  | 
val ((ps, prop'), ctxt') = Variable.focus NONE prop ctxt;  | 
| 60448 | 238  | 
fun fix (x, T) = (Binding.name (Variable.revert_fixed ctxt' x), SOME T, NoSyn);  | 
| 
42495
 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 
wenzelm 
parents: 
42494 
diff
changeset
 | 
239  | 
val xs = map (fix o #2) ps;  | 
| 
 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 
wenzelm 
parents: 
42494 
diff
changeset
 | 
240  | 
val As = Logic.strip_imp_prems prop';  | 
| 
41581
 
72a02e3dec7e
clarified pretty_statement: more robust treatment of fixes and conclusion of elimination (e.g. for classical rule);
 
wenzelm 
parents: 
41425 
diff
changeset
 | 
241  | 
in ((Binding.empty, (xs, As)), ctxt') end;  | 
| 19267 | 242  | 
|
243  | 
in  | 
|
244  | 
||
245  | 
fun pretty_statement ctxt kind raw_th =  | 
|
246  | 
let  | 
|
| 59970 | 247  | 
val (th, is_elim) = standard_elim ctxt (Raw_Simplifier.norm_hhf ctxt raw_th);  | 
| 
41581
 
72a02e3dec7e
clarified pretty_statement: more robust treatment of fixes and conclusion of elimination (e.g. for classical rule);
 
wenzelm 
parents: 
41425 
diff
changeset
 | 
248  | 
val ((_, [th']), ctxt') = Variable.import true [th] (Variable.set_body true ctxt);  | 
| 20150 | 249  | 
val prop = Thm.prop_of th';  | 
250  | 
val (prems, concl) = Logic.strip_horn prop;  | 
|
| 59970 | 251  | 
val concl_term = Object_Logic.drop_judgment ctxt concl;  | 
| 19267 | 252  | 
|
| 
62681
 
45b8dd2d3827
more accurate fixes (e.g. for notE, FalseE), amending baa589c574ff;
 
wenzelm 
parents: 
62680 
diff
changeset
 | 
253  | 
val (assumes, cases) =  | 
| 67522 | 254  | 
chop_suffix (fn prem => is_elim andalso concl aconv Logic.strip_assums_concl prem) prems;  | 
| 
62681
 
45b8dd2d3827
more accurate fixes (e.g. for notE, FalseE), amending baa589c574ff;
 
wenzelm 
parents: 
62680 
diff
changeset
 | 
255  | 
val is_thesis = if null cases then K false else fn v => v aconv concl_term;  | 
| 
 
45b8dd2d3827
more accurate fixes (e.g. for notE, FalseE), amending baa589c574ff;
 
wenzelm 
parents: 
62680 
diff
changeset
 | 
256  | 
val fixes =  | 
| 
 
45b8dd2d3827
more accurate fixes (e.g. for notE, FalseE), amending baa589c574ff;
 
wenzelm 
parents: 
62680 
diff
changeset
 | 
257  | 
rev (fold_aterms (fn v as Free (x, T) =>  | 
| 
 
45b8dd2d3827
more accurate fixes (e.g. for notE, FalseE), amending baa589c574ff;
 
wenzelm 
parents: 
62680 
diff
changeset
 | 
258  | 
if Variable.is_newly_fixed ctxt' ctxt x andalso not (is_thesis v)  | 
| 
 
45b8dd2d3827
more accurate fixes (e.g. for notE, FalseE), amending baa589c574ff;
 
wenzelm 
parents: 
62680 
diff
changeset
 | 
259  | 
then insert (op =) (Variable.revert_fixed ctxt' x, T) else I | _ => I) prop []);  | 
| 19267 | 260  | 
in  | 
| 28965 | 261  | 
pretty_ctxt ctxt' (Fixes (map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) fixes)) @  | 
| 63352 | 262  | 
pretty_ctxt ctxt' (Assumes (map (fn t => (Binding.empty_atts, [(t, [])])) assumes)) @  | 
263  | 
(if null cases then pretty_stmt ctxt' (Shows [(Binding.empty_atts, [(concl, [])])])  | 
|
| 
26716
 
8690e75e1395
print_statement: reset body mode, i.e. invent global frees (no need for revert_skolem);
 
wenzelm 
parents: 
26628 
diff
changeset
 | 
264  | 
else  | 
| 
42495
 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 
wenzelm 
parents: 
42494 
diff
changeset
 | 
265  | 
let val (clauses, ctxt'') = fold_map obtain cases ctxt'  | 
| 
26716
 
8690e75e1395
print_statement: reset body mode, i.e. invent global frees (no need for revert_skolem);
 
wenzelm 
parents: 
26628 
diff
changeset
 | 
266  | 
in pretty_stmt ctxt'' (Obtains clauses) end)  | 
| 
64398
 
5076725247fa
more robust printing of names in the context of outer syntax;
 
wenzelm 
parents: 
63395 
diff
changeset
 | 
267  | 
end |> thm_name ctxt kind raw_th;  | 
| 19267 | 268  | 
|
| 
18140
 
691c64d615a5
Explicit data structures for some Isar language elements.
 
wenzelm 
parents:  
diff
changeset
 | 
269  | 
end;  | 
| 19267 | 270  | 
|
| 19777 | 271  | 
|
272  | 
||
273  | 
(** logical operations **)  | 
|
274  | 
||
275  | 
(* witnesses -- hypotheses as protected facts *)  | 
|
276  | 
||
277  | 
datatype witness = Witness of term * thm;  | 
|
278  | 
||
| 29578 | 279  | 
val mark_witness = Logic.protect;  | 
280  | 
fun witness_prop (Witness (t, _)) = t;  | 
|
| 
77863
 
760515c45864
revert b43ee37926a9 due to problems with AFP/PAPP_Impossibility;
 
wenzelm 
parents: 
77808 
diff
changeset
 | 
281  | 
fun witness_hyps (Witness (_, th)) = Thm.hyps_of th;  | 
| 19777 | 282  | 
fun map_witness f (Witness witn) = Witness (f witn);  | 
283  | 
||
| 45290 | 284  | 
fun transform_witness phi = map_witness (fn (t, th) => (Morphism.term phi t, Morphism.thm phi th));  | 
| 21481 | 285  | 
|
| 20058 | 286  | 
fun prove_witness ctxt t tac =  | 
| 52732 | 287  | 
Witness (t,  | 
| 
71017
 
c85efa2be619
expose derivations more thoroughly, notably for locale/class reasoning;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
288  | 
Goal.prove ctxt [] [] (mark_witness t)  | 
| 
 
c85efa2be619
expose derivations more thoroughly, notably for locale/class reasoning;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
289  | 
(fn _ => resolve_tac ctxt [Drule.protectI] 1 THEN tac)  | 
| 78041 | 290  | 
|> Thm.close_derivation \<^here>  | 
291  | 
|> Thm.trim_context);  | 
|
| 19777 | 292  | 
|
| 29603 | 293  | 
|
| 29578 | 294  | 
local  | 
295  | 
||
296  | 
val refine_witness =  | 
|
| 70520 | 297  | 
Proof.refine_singleton (Method.Basic (fn ctxt => CONTEXT_TACTIC o  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59385 
diff
changeset
 | 
298  | 
K (ALLGOALS (CONJUNCTS (ALLGOALS (CONJUNCTS (TRYALL (resolve_tac ctxt [Drule.protectI]))))))));  | 
| 25624 | 299  | 
|
| 29578 | 300  | 
fun gen_witness_proof proof after_qed wit_propss eq_props =  | 
301  | 
let  | 
|
| 46896 | 302  | 
val propss =  | 
303  | 
(map o map) (fn prop => (mark_witness prop, [])) wit_propss @  | 
|
304  | 
[map (rpair []) eq_props];  | 
|
| 29578 | 305  | 
fun after_qed' thmss =  | 
| 78041 | 306  | 
let  | 
307  | 
val (wits, eqs) =  | 
|
308  | 
split_last ((map o map) (Thm.close_derivation \<^here> #> Thm.trim_context) thmss);  | 
|
| 29578 | 309  | 
in after_qed ((map2 o map2) (curry Witness) wit_propss wits) eqs end;  | 
| 
61841
 
4d3527b94f2a
more general types Proof.method / context_tactic;
 
wenzelm 
parents: 
61814 
diff
changeset
 | 
310  | 
in proof after_qed' propss #> refine_witness end;  | 
| 29578 | 311  | 
|
| 
62680
 
646b84666a56
eliminated unused argument (see also 58110c1e02bc);
 
wenzelm 
parents: 
62094 
diff
changeset
 | 
312  | 
fun proof_local cmd goal_ctxt after_qed propp =  | 
| 
60415
 
9d37b2330ee3
clarified local after_qed: result is not exported yet;
 
wenzelm 
parents: 
60414 
diff
changeset
 | 
313  | 
let  | 
| 
 
9d37b2330ee3
clarified local after_qed: result is not exported yet;
 
wenzelm 
parents: 
60414 
diff
changeset
 | 
314  | 
fun after_qed' (result_ctxt, results) state' =  | 
| 
 
9d37b2330ee3
clarified local after_qed: result is not exported yet;
 
wenzelm 
parents: 
60414 
diff
changeset
 | 
315  | 
after_qed (burrow (Proof_Context.export result_ctxt (Proof.context_of state')) results) state';  | 
| 
 
9d37b2330ee3
clarified local after_qed: result is not exported yet;
 
wenzelm 
parents: 
60414 
diff
changeset
 | 
316  | 
in  | 
| 
 
9d37b2330ee3
clarified local after_qed: result is not exported yet;
 
wenzelm 
parents: 
60414 
diff
changeset
 | 
317  | 
Proof.map_context (K goal_ctxt) #>  | 
| 
60555
 
51a6997b1384
support 'when' statement, which corresponds to 'presume';
 
wenzelm 
parents: 
60461 
diff
changeset
 | 
318  | 
Proof.internal_goal (K (K ())) (Proof_Context.get_mode goal_ctxt) true cmd  | 
| 63352 | 319  | 
NONE after_qed' [] [] (map (pair Binding.empty_atts) propp) #> snd  | 
| 
60415
 
9d37b2330ee3
clarified local after_qed: result is not exported yet;
 
wenzelm 
parents: 
60414 
diff
changeset
 | 
320  | 
end;  | 
| 41425 | 321  | 
|
| 29578 | 322  | 
in  | 
323  | 
||
324  | 
fun witness_proof after_qed wit_propss =  | 
|
| 
36323
 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 
wenzelm 
parents: 
35767 
diff
changeset
 | 
325  | 
gen_witness_proof (Proof.theorem NONE) (fn wits => fn _ => after_qed wits)  | 
| 29578 | 326  | 
wit_propss [];  | 
327  | 
||
| 
36323
 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 
wenzelm 
parents: 
35767 
diff
changeset
 | 
328  | 
val witness_proof_eqs = gen_witness_proof (Proof.theorem NONE);  | 
| 29578 | 329  | 
|
| 
62680
 
646b84666a56
eliminated unused argument (see also 58110c1e02bc);
 
wenzelm 
parents: 
62094 
diff
changeset
 | 
330  | 
fun witness_local_proof after_qed cmd wit_propss goal_ctxt =  | 
| 
 
646b84666a56
eliminated unused argument (see also 58110c1e02bc);
 
wenzelm 
parents: 
62094 
diff
changeset
 | 
331  | 
gen_witness_proof (proof_local cmd goal_ctxt)  | 
| 29578 | 332  | 
(fn wits => fn _ => after_qed wits) wit_propss [];  | 
333  | 
||
| 
62680
 
646b84666a56
eliminated unused argument (see also 58110c1e02bc);
 
wenzelm 
parents: 
62094 
diff
changeset
 | 
334  | 
fun witness_local_proof_eqs after_qed cmd wit_propss eq_props goal_ctxt =  | 
| 
 
646b84666a56
eliminated unused argument (see also 58110c1e02bc);
 
wenzelm 
parents: 
62094 
diff
changeset
 | 
335  | 
gen_witness_proof (proof_local cmd goal_ctxt) after_qed wit_propss eq_props;  | 
| 41425 | 336  | 
|
| 29603 | 337  | 
end;  | 
338  | 
||
| 
54883
 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 
wenzelm 
parents: 
54742 
diff
changeset
 | 
339  | 
fun conclude_witness ctxt (Witness (_, th)) =  | 
| 78041 | 340  | 
Goal.conclude (Thm.transfer' ctxt th)  | 
| 
71017
 
c85efa2be619
expose derivations more thoroughly, notably for locale/class reasoning;
 
wenzelm 
parents: 
70520 
diff
changeset
 | 
341  | 
|> Raw_Simplifier.norm_hhf_protect ctxt  | 
| 
71018
 
d32ed8927a42
more robust expose_proofs corresponding to register_proofs/consolidate_theory;
 
wenzelm 
parents: 
71017 
diff
changeset
 | 
342  | 
|> Thm.close_derivation \<^here>;  | 
| 19777 | 343  | 
|
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22568 
diff
changeset
 | 
344  | 
fun pretty_witness ctxt witn =  | 
| 24920 | 345  | 
let val prt_term = Pretty.quote o Syntax.pretty_term ctxt in  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22568 
diff
changeset
 | 
346  | 
Pretty.block (prt_term (witness_prop witn) ::  | 
| 
39166
 
19efc2af3e6c
turned show_hyps and show_tags into proper configuration option;
 
wenzelm 
parents: 
38709 
diff
changeset
 | 
347  | 
(if Config.get ctxt show_hyps then [Pretty.brk 2, Pretty.list "[" "]"  | 
| 
22658
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22568 
diff
changeset
 | 
348  | 
(map prt_term (witness_hyps witn))] else []))  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22568 
diff
changeset
 | 
349  | 
end;  | 
| 
 
263d42253f53
Experimental interpretation code for definitions.
 
ballarin 
parents: 
22568 
diff
changeset
 | 
350  | 
|
| 19777 | 351  | 
|
| 
67699
 
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
352  | 
(* instantiate frees, with beta normalization *)  | 
| 19777 | 353  | 
|
| 
67699
 
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
354  | 
fun instantiate_normalize_morphism insts =  | 
| 
 
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
355  | 
Morphism.instantiate_frees_morphism insts $>  | 
| 
 
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
356  | 
Morphism.term_morphism "beta_norm" Envir.beta_norm $>  | 
| 
 
8e4ff46f807d
more standard instantiate operations, where substitutions are already certified and stored with trimmed theory certificate;
 
wenzelm 
parents: 
67671 
diff
changeset
 | 
357  | 
Morphism.thm_morphism "beta_conversion" (Conv.fconv_rule (Thm.beta_conversion true));  | 
| 19777 | 358  | 
|
359  | 
||
360  | 
(* satisfy hypotheses *)  | 
|
361  | 
||
| 
67701
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
362  | 
local  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
363  | 
|
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
364  | 
val norm_term = Envir.beta_eta_contract;  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
365  | 
val norm_conv = Drule.beta_eta_conversion;  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
366  | 
val norm_cterm = Thm.rhs_of o norm_conv;  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
367  | 
|
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
368  | 
fun find_witness witns hyp =  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
369  | 
(case find_first (fn Witness (t, _) => hyp aconv t) witns of  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
370  | 
NONE =>  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
371  | 
let val hyp' = norm_term hyp  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
372  | 
in find_first (fn Witness (t, _) => hyp' aconv norm_term t) witns end  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
373  | 
| some => some);  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
374  | 
|
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
375  | 
fun compose_witness (Witness (_, th)) r =  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
376  | 
let  | 
| 78041 | 377  | 
val th' = Goal.conclude (Thm.transfer (Thm.theory_of_thm r) th);  | 
| 
67701
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
378  | 
val A = Thm.cprem_of r 1;  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
379  | 
in  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
380  | 
Thm.implies_elim  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
381  | 
(Conv.gconv_rule norm_conv 1 r)  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
382  | 
(Conv.fconv_rule norm_conv  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
383  | 
(Thm.instantiate (Thm.match (apply2 norm_cterm (Thm.cprop_of th', A))) th'))  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
384  | 
end;  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
385  | 
|
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
386  | 
in  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
387  | 
|
| 
45346
 
439101d8eeec
some performance tuning via Term_Subst/Same.operation;
 
wenzelm 
parents: 
45345 
diff
changeset
 | 
388  | 
fun satisfy_thm witns thm =  | 
| 
67701
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
389  | 
(Thm.chyps_of thm, thm) |-> fold (fn hyp =>  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
390  | 
(case find_witness witns (Thm.term_of hyp) of  | 
| 19777 | 391  | 
NONE => I  | 
| 
67701
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
392  | 
| SOME w => Thm.implies_intr hyp #> compose_witness w));  | 
| 19777 | 393  | 
|
| 54740 | 394  | 
val satisfy_morphism = Morphism.thm_morphism "Element.satisfy" o satisfy_thm;  | 
| 20264 | 395  | 
|
| 
67701
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
396  | 
end;  | 
| 
 
454dfdaf021d
more robust normalization of witness, e.g. relevant for proofs of transfer_existence_ivl0, transfer_flow0 in AFP/thys/Ordinary_Differential_Equations/Numerics/Transfer_Euclidean_Space_Vector.thy;
 
wenzelm 
parents: 
67699 
diff
changeset
 | 
397  | 
|
| 20264 | 398  | 
|
| 29525 | 399  | 
(* rewriting with equalities *)  | 
400  | 
||
| 
78453
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
401  | 
fun decomp_simp ctxt prop =  | 
| 
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
402  | 
let  | 
| 
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
403  | 
val _ = Logic.no_prems prop orelse  | 
| 
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
404  | 
      error ("Bad conditional rewrite rule " ^ Syntax.string_of_term ctxt prop);
 | 
| 
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
405  | 
in  | 
| 
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
406  | 
Logic.dest_equals prop handle TERM _ =>  | 
| 
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
407  | 
      error ("Rewrite rule not a meta-equality " ^ Syntax.string_of_term ctxt prop)
 | 
| 
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
408  | 
end;  | 
| 
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
409  | 
|
| 67740 | 410  | 
(* for activating declarations only *)  | 
| 
78453
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
411  | 
fun eq_term_morphism _ [] = NONE  | 
| 
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
412  | 
| eq_term_morphism ctxt0 props =  | 
| 67740 | 413  | 
let  | 
| 
78453
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
414  | 
val simps = map (decomp_simp ctxt0) props;  | 
| 
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
415  | 
|
| 
78065
 
11d6a1e62841
more careful treatment of set_context / reset_context for persistent morphisms;
 
wenzelm 
parents: 
78064 
diff
changeset
 | 
416  | 
fun rewrite_term thy =  | 
| 
 
11d6a1e62841
more careful treatment of set_context / reset_context for persistent morphisms;
 
wenzelm 
parents: 
78064 
diff
changeset
 | 
417  | 
let val ctxt = Proof_Context.init_global thy  | 
| 
78453
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
418  | 
in Pattern.rewrite_term thy simps [] end;  | 
| 67740 | 419  | 
val phi =  | 
420  | 
Morphism.morphism "Element.eq_term_morphism"  | 
|
421  | 
           {binding = [],
 | 
|
422  | 
typ = [],  | 
|
| 
78065
 
11d6a1e62841
more careful treatment of set_context / reset_context for persistent morphisms;
 
wenzelm 
parents: 
78064 
diff
changeset
 | 
423  | 
term = [rewrite_term o Morphism.the_theory],  | 
| 
78062
 
edb195122938
support for context within morphism (for background theory);
 
wenzelm 
parents: 
78041 
diff
changeset
 | 
424  | 
fact = [fn _ => fn _ => error "Illegal application of Element.eq_term_morphism"]};  | 
| 67740 | 425  | 
in SOME phi end;  | 
426  | 
||
| 
78453
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
427  | 
fun eq_morphism _ [] = NONE  | 
| 
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
428  | 
| eq_morphism ctxt0 thms =  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54740 
diff
changeset
 | 
429  | 
let  | 
| 
78453
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
430  | 
val simpset = Raw_Simplifier.simpset_of (Raw_Simplifier.init_simpset thms ctxt0);  | 
| 
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
431  | 
val simps = map (decomp_simp ctxt0 o Thm.full_prop_of) (Raw_Simplifier.dest_simps simpset);  | 
| 
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
432  | 
|
| 
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
433  | 
fun rewrite_term thy = Pattern.rewrite_term thy simps [];  | 
| 
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
434  | 
val rewrite =  | 
| 
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
435  | 
Proof_Context.init_global #>  | 
| 
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
436  | 
Raw_Simplifier.put_simpset simpset #>  | 
| 
 
3fdf3c5cfa9d
performance tuning: prefer static simpset within functional closure of morphism (with notable impact on specifications "in" class, e.g. AFP/No_FTL_observers);
 
wenzelm 
parents: 
78065 
diff
changeset
 | 
437  | 
Raw_Simplifier.rewrite0_rule;  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54740 
diff
changeset
 | 
438  | 
val phi =  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54740 
diff
changeset
 | 
439  | 
Morphism.morphism "Element.eq_morphism"  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54740 
diff
changeset
 | 
440  | 
           {binding = [],
 | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54740 
diff
changeset
 | 
441  | 
typ = [],  | 
| 
78065
 
11d6a1e62841
more careful treatment of set_context / reset_context for persistent morphisms;
 
wenzelm 
parents: 
78064 
diff
changeset
 | 
442  | 
term = [rewrite_term o Morphism.the_theory],  | 
| 
 
11d6a1e62841
more careful treatment of set_context / reset_context for persistent morphisms;
 
wenzelm 
parents: 
78064 
diff
changeset
 | 
443  | 
fact = [map o rewrite o Morphism.the_theory]};  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54740 
diff
changeset
 | 
444  | 
in SOME phi end;  | 
| 29525 | 445  | 
|
446  | 
||
| 29218 | 447  | 
|
| 
30775
 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
448  | 
(** activate in context **)  | 
| 28832 | 449  | 
|
| 
30775
 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
450  | 
(* init *)  | 
| 28832 | 451  | 
|
| 
57864
 
7cf01ece66e4
clarified Element.init vs. Element.init' -- the latter also avoids redundant warnings due to declatations when preparing locale expressions / interpretations;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
452  | 
fun init (Fixes fixes) = Context.map_proof (Proof_Context.add_fixes fixes #> #2)  | 
| 
 
7cf01ece66e4
clarified Element.init vs. Element.init' -- the latter also avoids redundant warnings due to declatations when preparing locale expressions / interpretations;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
453  | 
| init (Constrains _) = I  | 
| 
 
7cf01ece66e4
clarified Element.init vs. Element.init' -- the latter also avoids redundant warnings due to declatations when preparing locale expressions / interpretations;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
454  | 
| init (Assumes asms) = Context.map_proof (fn ctxt =>  | 
| 28832 | 455  | 
let  | 
| 47815 | 456  | 
val asms' = Attrib.map_specs (map (Attrib.attribute ctxt)) asms;  | 
| 
30775
 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
457  | 
val (_, ctxt') = ctxt  | 
| 70308 | 458  | 
|> fold Proof_Context.augment (maps (map #1 o #2) asms')  | 
| 60377 | 459  | 
|> Proof_Context.add_assms Assumption.assume_export asms';  | 
| 
30775
 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
460  | 
in ctxt' end)  | 
| 
57864
 
7cf01ece66e4
clarified Element.init vs. Element.init' -- the latter also avoids redundant warnings due to declatations when preparing locale expressions / interpretations;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
461  | 
| init (Defines defs) = Context.map_proof (fn ctxt =>  | 
| 28832 | 462  | 
let  | 
| 47815 | 463  | 
val defs' = Attrib.map_specs (map (Attrib.attribute ctxt)) defs;  | 
| 
49750
 
444cfaa331c9
clarified Element.init vs. Element.activate: refrain from hard-wiring Thm.def_binding_optional to avoid duplicate facts;
 
wenzelm 
parents: 
47815 
diff
changeset
 | 
464  | 
val asms = defs' |> map (fn (b, (t, ps)) =>  | 
| 63395 | 465  | 
let val (_, t') = Local_Defs.cert_def ctxt (K []) t (* FIXME adapt ps? *)  | 
| 
49750
 
444cfaa331c9
clarified Element.init vs. Element.activate: refrain from hard-wiring Thm.def_binding_optional to avoid duplicate facts;
 
wenzelm 
parents: 
47815 
diff
changeset
 | 
466  | 
in (t', (b, [(t', ps)])) end);  | 
| 
30775
 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
467  | 
val (_, ctxt') = ctxt  | 
| 70308 | 468  | 
|> fold Proof_Context.augment (map #1 asms)  | 
| 60377 | 469  | 
|> Proof_Context.add_assms Local_Defs.def_export (map #2 asms);  | 
| 
30775
 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
470  | 
in ctxt' end)  | 
| 
67671
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
471  | 
| init (Notes (kind, facts)) = Attrib.generic_notes kind facts #> #2  | 
| 
 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 
wenzelm 
parents: 
67667 
diff
changeset
 | 
472  | 
| init (Lazy_Notes (kind, ths)) = Attrib.lazy_notes kind ths;  | 
| 
54993
 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 
wenzelm 
parents: 
54883 
diff
changeset
 | 
473  | 
|
| 
30775
 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
474  | 
|
| 
 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
475  | 
(* activate *)  | 
| 
 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 
wenzelm 
parents: 
30763 
diff
changeset
 | 
476  | 
|
| 
30777
 
9960ff945c52
simplified Element.activate(_i): singleton version;
 
wenzelm 
parents: 
30775 
diff
changeset
 | 
477  | 
fun activate_i elem ctxt =  | 
| 28832 | 478  | 
let  | 
| 
49750
 
444cfaa331c9
clarified Element.init vs. Element.activate: refrain from hard-wiring Thm.def_binding_optional to avoid duplicate facts;
 
wenzelm 
parents: 
47815 
diff
changeset
 | 
479  | 
val elem' =  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
480  | 
(case (map_ctxt_attrib o map) Token.init_assignable elem of  | 
| 
49750
 
444cfaa331c9
clarified Element.init vs. Element.activate: refrain from hard-wiring Thm.def_binding_optional to avoid duplicate facts;
 
wenzelm 
parents: 
47815 
diff
changeset
 | 
481  | 
Defines defs =>  | 
| 
 
444cfaa331c9
clarified Element.init vs. Element.activate: refrain from hard-wiring Thm.def_binding_optional to avoid duplicate facts;
 
wenzelm 
parents: 
47815 
diff
changeset
 | 
482  | 
Defines (defs |> map (fn ((a, atts), (t, ps)) =>  | 
| 63395 | 483  | 
((Thm.def_binding_optional  | 
484  | 
(Binding.name (#1 (#1 (Local_Defs.cert_def ctxt (K []) t)))) a, atts),  | 
|
| 
49750
 
444cfaa331c9
clarified Element.init vs. Element.activate: refrain from hard-wiring Thm.def_binding_optional to avoid duplicate facts;
 
wenzelm 
parents: 
47815 
diff
changeset
 | 
485  | 
(t, ps))))  | 
| 
 
444cfaa331c9
clarified Element.init vs. Element.activate: refrain from hard-wiring Thm.def_binding_optional to avoid duplicate facts;
 
wenzelm 
parents: 
47815 
diff
changeset
 | 
486  | 
| e => e);  | 
| 
30777
 
9960ff945c52
simplified Element.activate(_i): singleton version;
 
wenzelm 
parents: 
30775 
diff
changeset
 | 
487  | 
val ctxt' = Context.proof_map (init elem') ctxt;  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
488  | 
in ((map_ctxt_attrib o map) Token.closure elem', ctxt') end;  | 
| 28832 | 489  | 
|
| 
30777
 
9960ff945c52
simplified Element.activate(_i): singleton version;
 
wenzelm 
parents: 
30775 
diff
changeset
 | 
490  | 
fun activate raw_elem ctxt =  | 
| 
 
9960ff945c52
simplified Element.activate(_i): singleton version;
 
wenzelm 
parents: 
30775 
diff
changeset
 | 
491  | 
let val elem = raw_elem |> map_ctxt  | 
| 
43842
 
f035d867fb41
Element.activate: leave check of binding where actually applied to the context -- allow internal qualifications, or non-identifier fact names like "assumes *: A" (see also 1183951365de);
 
wenzelm 
parents: 
43837 
diff
changeset
 | 
492  | 
   {binding = I,
 | 
| 29603 | 493  | 
typ = I,  | 
494  | 
term = I,  | 
|
495  | 
pattern = I,  | 
|
| 42360 | 496  | 
fact = Proof_Context.get_fact ctxt,  | 
| 
55997
 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 
wenzelm 
parents: 
55914 
diff
changeset
 | 
497  | 
attrib = Attrib.check_src ctxt}  | 
| 
30777
 
9960ff945c52
simplified Element.activate(_i): singleton version;
 
wenzelm 
parents: 
30775 
diff
changeset
 | 
498  | 
in activate_i elem ctxt end;  | 
| 28832 | 499  | 
|
| 19267 | 500  | 
end;  |