| author | Fabian Huch <huch@in.tum.de> | 
| Wed, 03 Jul 2024 21:11:24 +0200 | |
| changeset 80497 | bd00bdf39c86 | 
| parent 78453 | 3fdf3c5cfa9d | 
| child 81533 | fb49af893986 | 
| permissions | -rw-r--r-- | 
| 61669 | 1 | (* Title: Pure/Isar/interpretation.ML | 
| 2 | Author: Clemens Ballarin, TU Muenchen | |
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 3 | Author: Florian Haftmann, TU Muenchen | 
| 61669 | 4 | |
| 5 | Locale interpretation. | |
| 6 | *) | |
| 7 | ||
| 8 | signature INTERPRETATION = | |
| 9 | sig | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 10 | type 'a defines = (Attrib.binding * ((binding * mixfix) * 'a)) list | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 11 | |
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 12 | (*interpretation in proofs*) | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 13 | val interpret: Expression.expression_i -> Proof.state -> Proof.state | 
| 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 14 | val interpret_cmd: Expression.expression -> Proof.state -> Proof.state | 
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 15 | |
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 16 | (*interpretation in local theories*) | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 17 | val interpretation: Expression.expression_i -> local_theory -> Proof.state | 
| 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 18 | val interpretation_cmd: Expression.expression -> local_theory -> Proof.state | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 19 | |
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 20 | (*interpretation into global theories*) | 
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 21 | val global_interpretation: Expression.expression_i -> | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 22 | term defines -> local_theory -> Proof.state | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 23 | val global_interpretation_cmd: Expression.expression -> | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 24 | string defines -> local_theory -> Proof.state | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 25 | |
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 26 | (*interpretation between locales*) | 
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 27 | val sublocale: Expression.expression_i -> | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 28 | term defines -> local_theory -> Proof.state | 
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 29 | val sublocale_cmd: Expression.expression -> | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 30 | string defines -> local_theory -> Proof.state | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 31 | val global_sublocale: string -> Expression.expression_i -> | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 32 | term defines -> theory -> Proof.state | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 33 | val global_sublocale_cmd: xstring * Position.T -> Expression.expression -> | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 34 | string defines -> theory -> Proof.state | 
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 35 | |
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 36 | (*mixed Isar interface*) | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 37 | val isar_interpretation: Expression.expression_i -> local_theory -> Proof.state | 
| 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 38 | val isar_interpretation_cmd: Expression.expression -> local_theory -> Proof.state | 
| 61669 | 39 | end; | 
| 40 | ||
| 41 | structure Interpretation : INTERPRETATION = | |
| 42 | struct | |
| 43 | ||
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 44 | (** common interpretation machinery **) | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 45 | |
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 46 | type 'a defines = (Attrib.binding * ((binding * mixfix) * 'a)) list | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 47 | |
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 48 | (* reading of locale expressions with rewrite morphisms *) | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 49 | |
| 61672 
87203a0f0041
tuned -- share implementations as far as appropriate
 haftmann parents: 
61670diff
changeset | 50 | local | 
| 61669 | 51 | |
| 67450 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 52 | fun augment_with_def prep_term ((name, atts), ((b, mx), raw_rhs)) lthy = | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 53 | let | 
| 61775 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 54 | val rhs = prep_term lthy raw_rhs; | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 55 | val lthy' = Variable.declare_term rhs lthy; | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 56 | val ((_, (_, def)), lthy'') = | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 57 | Local_Theory.define ((b, mx), ((Thm.def_binding_optional b name, atts), rhs)) lthy'; | 
| 69699 
82f57315cade
dedicated combinator for declarations nested in a local theory block
 haftmann parents: 
69058diff
changeset | 58 | in (Thm.symmetric def, lthy'') end; | 
| 61669 | 59 | |
| 69699 
82f57315cade
dedicated combinator for declarations nested in a local theory block
 haftmann parents: 
69058diff
changeset | 60 | fun augment_with_defs _ [] _ = pair [] | 
| 61691 
91854ba66adb
clarified contexts by factoring out reading and definition of mixins
 haftmann parents: 
61684diff
changeset | 61 | (*quasi-inhomogeneous type: definitions demand local theory rather than bare proof context*) | 
| 69699 
82f57315cade
dedicated combinator for declarations nested in a local theory block
 haftmann parents: 
69058diff
changeset | 62 | | augment_with_defs prep_term raw_defs deps = | 
| 72451 
e51f1733618d
replaced combinators by more conventional nesting pattern
 haftmann parents: 
70314diff
changeset | 63 | Local_Theory.begin_nested | 
| 
e51f1733618d
replaced combinators by more conventional nesting pattern
 haftmann parents: 
70314diff
changeset | 64 | #> snd | 
| 
e51f1733618d
replaced combinators by more conventional nesting pattern
 haftmann parents: 
70314diff
changeset | 65 | #> fold Locale.activate_declarations deps | 
| 
e51f1733618d
replaced combinators by more conventional nesting pattern
 haftmann parents: 
70314diff
changeset | 66 | #> fold_map (augment_with_def prep_term) raw_defs | 
| 
e51f1733618d
replaced combinators by more conventional nesting pattern
 haftmann parents: 
70314diff
changeset | 67 | #> Local_Theory.end_nested_result Morphism.fact; | 
| 61775 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 68 | |
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 69 | fun prep_interpretation prep_expr prep_term | 
| 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 70 | expression raw_defs initial_ctxt = | 
| 61691 
91854ba66adb
clarified contexts by factoring out reading and definition of mixins
 haftmann parents: 
61684diff
changeset | 71 | let | 
| 67450 
b0ae74b86ef3
Experimental support for rewrite morphisms in locale instances.
 ballarin parents: 
66334diff
changeset | 72 | val ((propss, eq_propss, deps, eqnss, export), expr_ctxt) = prep_expr expression initial_ctxt; | 
| 61775 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 73 | val (def_eqns, def_ctxt) = | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 74 | augment_with_defs prep_term raw_defs deps expr_ctxt; | 
| 70314 
6d6839a948cf
proper Proof_Context.export_morphism corresponding to Proof_Context.augment (see 7f568724d67e);
 wenzelm parents: 
69699diff
changeset | 75 | val export' = Proof_Context.export_morphism def_ctxt expr_ctxt; | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 76 | in (((propss, eq_propss, deps, eqnss, export, export'), def_eqns), def_ctxt) end; | 
| 61669 | 77 | |
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 78 | in | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 79 | |
| 61698 | 80 | fun cert_interpretation expression = | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 81 | prep_interpretation Expression.cert_goal_expression Syntax.check_term expression; | 
| 61672 
87203a0f0041
tuned -- share implementations as far as appropriate
 haftmann parents: 
61670diff
changeset | 82 | |
| 61698 | 83 | fun read_interpretation expression = | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 84 | prep_interpretation Expression.read_goal_expression Syntax.read_term expression; | 
| 61669 | 85 | |
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 86 | end; | 
| 61669 | 87 | |
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 88 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 89 | (* interpretation machinery *) | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 90 | |
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 91 | local | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 92 | |
| 68854 | 93 | fun abs_def_rule eqns ctxt = | 
| 63381 | 94 | (map (Local_Defs.abs_def_rule ctxt) (maps snd eqns), ctxt); | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 95 | |
| 68855 
59ce31e15c33
more careful treatment position: existing facts refer to interpretation command, future facts refer to themselves (see also 4270da306442);
 wenzelm parents: 
68854diff
changeset | 96 | fun note_eqns_register note add_registration | 
| 68849 | 97 | deps eqnss witss def_eqns thms export export' ctxt = | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 98 | let | 
| 68849 | 99 | val factss = thms | 
| 100 | |> unflat ((map o map) #1 eqnss) | |
| 78041 | 101 | |> map2 (map2 (fn b => fn eq => | 
| 102 | (b, [([Morphism.thm export (Thm.transfer' ctxt eq)], [])]))) ((map o map) #1 eqnss); | |
| 68849 | 103 | val (eqnss', ctxt') = | 
| 68854 | 104 | fold_map (fn facts => note Thm.theoremK facts #-> abs_def_rule) factss ctxt; | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 105 | val defs = (Binding.empty_atts, [(map (Morphism.thm (export' $> export)) def_eqns, [])]); | 
| 68854 | 106 | val (eqns', ctxt'') = ctxt' |> note Thm.theoremK [defs] |-> abs_def_rule; | 
| 78064 
4e865c45458b
clarified transfer / trim_context on persistent Token.source (e.g. attribute expressions): actually set/reset implicit context;
 wenzelm parents: 
78047diff
changeset | 107 | val transform_witness = Element.transform_witness (Morphism.set_trim_context' ctxt' export'); | 
| 68854 | 108 | val deps' = | 
| 109 | (deps ~~ witss) |> map (fn ((dep, morph), wits) => | |
| 78047 | 110 | (dep, morph $> Element.satisfy_morphism (map transform_witness wits))); | 
| 68854 | 111 | fun register (dep, eqns) ctxt = | 
| 68851 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 wenzelm parents: 
68850diff
changeset | 112 | ctxt |> add_registration | 
| 69058 | 113 |         {inst = dep,
 | 
| 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: 
78065diff
changeset | 114 | mixin = Option.map (rpair true) (Element.eq_morphism ctxt (eqns @ eqns')), | 
| 68851 
6c9825c1e26b
clarified signature: explicit type Locale.registration;
 wenzelm parents: 
68850diff
changeset | 115 | export = export}; | 
| 68854 | 116 | in ctxt'' |> fold register (deps' ~~ eqnss') end; | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 117 | |
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 118 | in | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 119 | |
| 63380 | 120 | fun generic_interpretation prep_interpretation setup_proof note add_registration | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 121 | expression raw_defs initial_ctxt = | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 122 | let | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 123 | val (((propss, eq_propss, deps, eqnss, export, export'), def_eqns), goal_ctxt) = | 
| 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 124 | prep_interpretation expression raw_defs initial_ctxt; | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 125 | fun after_qed witss eqns = | 
| 68855 
59ce31e15c33
more careful treatment position: existing facts refer to interpretation command, future facts refer to themselves (see also 4270da306442);
 wenzelm parents: 
68854diff
changeset | 126 | note_eqns_register note add_registration deps eqnss witss def_eqns eqns export export'; | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 127 | in setup_proof after_qed propss (flat eq_propss) goal_ctxt end; | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 128 | |
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 129 | end; | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 130 | |
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 131 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 132 | (** interfaces **) | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 133 | |
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 134 | (* interpretation in proofs *) | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 135 | |
| 61672 
87203a0f0041
tuned -- share implementations as far as appropriate
 haftmann parents: 
61670diff
changeset | 136 | local | 
| 
87203a0f0041
tuned -- share implementations as far as appropriate
 haftmann parents: 
61670diff
changeset | 137 | |
| 68850 | 138 | fun setup_proof state after_qed propss eqns goal_ctxt = | 
| 139 | Element.witness_local_proof_eqs | |
| 140 | (fn witss => fn eqns => Proof.map_context (after_qed witss eqns) #> Proof.reset_facts) | |
| 141 | "interpret" propss eqns goal_ctxt state; | |
| 142 | ||
| 73845 | 143 | fun add_registration_proof registration ctxt = ctxt | 
| 144 | |> Proof_Context.set_stmt false | |
| 145 | |> Context.proof_map (Locale.add_registration registration) | |
| 146 | |> Proof_Context.restore_stmt ctxt; | |
| 147 | ||
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 148 | fun gen_interpret prep_interpretation expression state = | 
| 68850 | 149 | Proof.assert_forward_or_chain state | 
| 150 | |> Proof.context_of | |
| 151 | |> generic_interpretation prep_interpretation (setup_proof state) | |
| 73845 | 152 | Attrib.local_notes add_registration_proof expression []; | 
| 61669 | 153 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 154 | in | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 155 | |
| 63379 | 156 | val interpret = gen_interpret cert_interpretation; | 
| 157 | val interpret_cmd = gen_interpret read_interpretation; | |
| 61669 | 158 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 159 | end; | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 160 | |
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 161 | |
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 162 | (* interpretation in local theories *) | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 163 | |
| 72953 | 164 | val add_registration_local_theory = | 
| 73845 | 165 | Named_Target.revoke_reinitializability oo Generic_Target.local_interpretation; | 
| 72953 | 166 | |
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 167 | fun interpretation expression = | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 168 | generic_interpretation cert_interpretation Element.witness_proof_eqs | 
| 72953 | 169 | Local_Theory.notes_kind add_registration_local_theory expression []; | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 170 | |
| 63380 | 171 | fun interpretation_cmd expression = | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 172 | generic_interpretation read_interpretation Element.witness_proof_eqs | 
| 72953 | 173 | Local_Theory.notes_kind add_registration_local_theory expression []; | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 174 | |
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 175 | |
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 176 | (* interpretation into global theories *) | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 177 | |
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 178 | fun global_interpretation expression = | 
| 63380 | 179 | generic_interpretation cert_interpretation Element.witness_proof_eqs | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 180 | Local_Theory.notes_kind Local_Theory.theory_registration expression; | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 181 | |
| 63380 | 182 | fun global_interpretation_cmd expression = | 
| 183 | generic_interpretation read_interpretation Element.witness_proof_eqs | |
| 184 | Local_Theory.notes_kind Local_Theory.theory_registration expression; | |
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 185 | |
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 186 | |
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 187 | (* interpretation between locales *) | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 188 | |
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 189 | fun sublocale expression = | 
| 63380 | 190 | generic_interpretation cert_interpretation Element.witness_proof_eqs | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 191 | Local_Theory.notes_kind Local_Theory.locale_dependency expression; | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 192 | |
| 63380 | 193 | fun sublocale_cmd expression = | 
| 194 | generic_interpretation read_interpretation Element.witness_proof_eqs | |
| 195 | Local_Theory.notes_kind Local_Theory.locale_dependency expression; | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 196 | |
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 197 | local | 
| 61669 | 198 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 199 | fun gen_global_sublocale prep_loc prep_interpretation | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 200 | raw_locale expression raw_defs thy = | 
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 201 | let | 
| 72536 
589645894305
bundle mixins for locale and class specifications
 haftmann parents: 
72505diff
changeset | 202 | val lthy = Named_Target.init [] (prep_loc thy raw_locale) thy; | 
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 203 | fun setup_proof after_qed = | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 204 | Element.witness_proof_eqs | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 205 | (fn wits => fn eqs => after_qed wits eqs #> Local_Theory.exit); | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 206 | in | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 207 | lthy |> | 
| 63380 | 208 | generic_interpretation prep_interpretation setup_proof | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 209 | Local_Theory.notes_kind Local_Theory.locale_dependency expression raw_defs | 
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 210 | end; | 
| 61669 | 211 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 212 | in | 
| 61669 | 213 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 214 | fun global_sublocale expression = | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 215 | gen_global_sublocale (K I) cert_interpretation expression; | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 216 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 217 | fun global_sublocale_cmd raw_expression = | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 218 | gen_global_sublocale Locale.check read_interpretation raw_expression; | 
| 61669 | 219 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 220 | end; | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 221 | |
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 222 | |
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 223 | (* mixed Isar interface *) | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 224 | |
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 225 | local | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 226 | |
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 227 | fun register_or_activate lthy = | 
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 228 | if Named_Target.is_theory lthy | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 229 | then Local_Theory.theory_registration | 
| 72953 | 230 | else add_registration_local_theory; | 
| 61669 | 231 | |
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 232 | fun gen_isar_interpretation prep_interpretation expression lthy = | 
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 233 | generic_interpretation prep_interpretation Element.witness_proof_eqs | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67740diff
changeset | 234 | Local_Theory.notes_kind (register_or_activate lthy) expression [] lthy; | 
| 61669 | 235 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 236 | in | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 237 | |
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 238 | fun isar_interpretation expression = | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 239 | gen_isar_interpretation cert_interpretation expression; | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 240 | fun isar_interpretation_cmd raw_expression = | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 241 | gen_isar_interpretation read_interpretation raw_expression; | 
| 61669 | 242 | |
| 243 | end; | |
| 244 | ||
| 245 | end; |