| author | wenzelm | 
| Mon, 07 Mar 2016 20:44:47 +0100 | |
| changeset 62548 | f8ebb715e06d | 
| parent 61890 | f6ded81f5690 | 
| child 62680 | 646b84666a56 | 
| 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 | type 'a rewrites = (Attrib.binding * 'a) list | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 12 | |
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 13 | (*interpretation in proofs*) | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 14 | val interpret: Expression.expression_i -> | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 15 | term rewrites -> bool -> Proof.state -> Proof.state | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 16 | val interpret_cmd: Expression.expression -> | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 17 | string rewrites -> bool -> Proof.state -> Proof.state | 
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 18 | |
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 19 | (*interpretation in local theories*) | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 20 | val interpretation: Expression.expression_i -> | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 21 | term rewrites -> local_theory -> Proof.state | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 22 | val interpretation_cmd: Expression.expression -> | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 23 | string rewrites -> local_theory -> Proof.state | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 24 | |
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 25 | (*interpretation into global theories*) | 
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 26 | val global_interpretation: Expression.expression_i -> | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 27 | term defines -> term rewrites -> local_theory -> Proof.state | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 28 | val global_interpretation_cmd: Expression.expression -> | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 29 | string defines -> string rewrites -> local_theory -> Proof.state | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 30 | |
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 31 | (*interpretation between locales*) | 
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 32 | val sublocale: Expression.expression_i -> | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 33 | term defines -> term rewrites -> local_theory -> Proof.state | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 34 | val sublocale_cmd: Expression.expression -> | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 35 | string defines -> string rewrites -> local_theory -> Proof.state | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 36 | val global_sublocale: string -> Expression.expression_i -> | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 37 | term defines -> term rewrites -> theory -> Proof.state | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 38 | val global_sublocale_cmd: xstring * Position.T -> Expression.expression -> | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 39 | string defines -> string rewrites -> theory -> Proof.state | 
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 40 | |
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 41 | (*mixed Isar interface*) | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 42 | val isar_interpretation: Expression.expression_i -> | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 43 | term rewrites -> local_theory -> Proof.state | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 44 | val isar_interpretation_cmd: Expression.expression -> | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 45 | string rewrites -> local_theory -> Proof.state | 
| 61669 | 46 | end; | 
| 47 | ||
| 48 | structure Interpretation : INTERPRETATION = | |
| 49 | struct | |
| 50 | ||
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 51 | (** common interpretation machinery **) | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 52 | |
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 53 | 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 | 54 | type 'a rewrites = (Attrib.binding * 'a) list | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 55 | |
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 56 | (* reading of locale expressions with rewrite morphisms *) | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 57 | |
| 61672 
87203a0f0041
tuned -- share implementations as far as appropriate
 haftmann parents: 
61670diff
changeset | 58 | local | 
| 61669 | 59 | |
| 61775 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 60 | fun augment_with_def prep_term deps ((name, atts), ((b, mx), raw_rhs)) lthy = | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 61 | let | 
| 61775 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 62 | val rhs = prep_term lthy raw_rhs; | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 63 | val lthy' = Variable.declare_term rhs lthy; | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 64 | val ((_, (_, def)), lthy'') = | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 65 | Local_Theory.define ((b, mx), ((Thm.def_binding_optional b name, atts), rhs)) lthy'; | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 66 | in (def, lthy'') end; | 
| 61669 | 67 | |
| 61775 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 68 | fun augment_with_defs prep_term [] deps ctxt = ([], ctxt) | 
| 61691 
91854ba66adb
clarified contexts by factoring out reading and definition of mixins
 haftmann parents: 
61684diff
changeset | 69 | (*quasi-inhomogeneous type: definitions demand local theory rather than bare proof context*) | 
| 61775 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 70 | | augment_with_defs prep_term raw_defs deps lthy = | 
| 61691 
91854ba66adb
clarified contexts by factoring out reading and definition of mixins
 haftmann parents: 
61684diff
changeset | 71 | let | 
| 61775 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 72 | val (_, inner_lthy) = | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 73 | Local_Theory.open_target lthy | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 74 | ||> fold Locale.activate_declarations deps; | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 75 | val (inner_defs, inner_lthy') = | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 76 | fold_map (augment_with_def prep_term deps) raw_defs inner_lthy; | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 77 | val lthy' = | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 78 | inner_lthy' | 
| 61708 
4de2380ae3ab
explicit nested local theory for definitions, however retaining arcane low-level fiddling with background theory
 haftmann parents: 
61698diff
changeset | 79 | |> Local_Theory.close_target; | 
| 
4de2380ae3ab
explicit nested local theory for definitions, however retaining arcane low-level fiddling with background theory
 haftmann parents: 
61698diff
changeset | 80 | val def_eqns = | 
| 61775 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 81 | map (singleton (Proof_Context.export inner_lthy' lthy') o Thm.symmetric) inner_defs | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 82 | in (def_eqns, lthy') end; | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 83 | |
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 84 | fun prep_eqns prep_props prep_attr [] deps ctxt = ([], []) | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 85 | | prep_eqns prep_props prep_attr raw_eqns deps ctxt = | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 86 | let | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 87 | val ctxt' = fold Locale.activate_declarations deps ctxt; | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 88 | val eqns = | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 89 | (Variable.export_terms ctxt' ctxt o prep_props ctxt' o map snd) raw_eqns; | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 90 | val attrss = map (apsnd (map (prep_attr ctxt)) o fst) raw_eqns; | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 91 | in (eqns, attrss) end; | 
| 61691 
91854ba66adb
clarified contexts by factoring out reading and definition of mixins
 haftmann parents: 
61684diff
changeset | 92 | |
| 61774 
029162bc9793
prefer conventional read/check distinction over manual check
 haftmann parents: 
61773diff
changeset | 93 | fun prep_interpretation prep_expr prep_term prep_props prep_attr | 
| 61691 
91854ba66adb
clarified contexts by factoring out reading and definition of mixins
 haftmann parents: 
61684diff
changeset | 94 | expression raw_defs raw_eqns initial_ctxt = | 
| 
91854ba66adb
clarified contexts by factoring out reading and definition of mixins
 haftmann parents: 
61684diff
changeset | 95 | let | 
| 
91854ba66adb
clarified contexts by factoring out reading and definition of mixins
 haftmann parents: 
61684diff
changeset | 96 | val ((propss, deps, 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 | 97 | val (def_eqns, def_ctxt) = | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 98 | augment_with_defs prep_term raw_defs deps expr_ctxt; | 
| 
ec11275fb263
alternating parsing and defining of rewrite definitions: formally correct treatment of polymorphism
 haftmann parents: 
61774diff
changeset | 99 | val (eqns, attrss) = prep_eqns prep_props prep_attr raw_eqns deps def_ctxt; | 
| 61672 
87203a0f0041
tuned -- share implementations as far as appropriate
 haftmann parents: 
61670diff
changeset | 100 | val goal_ctxt = fold Variable.auto_fixes eqns def_ctxt; | 
| 61772 
2f33f6cc964d
formally correct context for export, which got screwed up in 87203a0f0041
 haftmann parents: 
61771diff
changeset | 101 | val export' = Variable.export_morphism goal_ctxt expr_ctxt; | 
| 61672 
87203a0f0041
tuned -- share implementations as far as appropriate
 haftmann parents: 
61670diff
changeset | 102 | in (((propss, deps, export, export'), (def_eqns, eqns, attrss)), goal_ctxt) end; | 
| 61669 | 103 | |
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 104 | in | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 105 | |
| 61698 | 106 | fun cert_interpretation expression = | 
| 61774 
029162bc9793
prefer conventional read/check distinction over manual check
 haftmann parents: 
61773diff
changeset | 107 | prep_interpretation Expression.cert_goal_expression Syntax.check_term | 
| 
029162bc9793
prefer conventional read/check distinction over manual check
 haftmann parents: 
61773diff
changeset | 108 | Syntax.check_props (K I) expression; | 
| 61672 
87203a0f0041
tuned -- share implementations as far as appropriate
 haftmann parents: 
61670diff
changeset | 109 | |
| 61698 | 110 | fun read_interpretation expression = | 
| 61774 
029162bc9793
prefer conventional read/check distinction over manual check
 haftmann parents: 
61773diff
changeset | 111 | prep_interpretation Expression.read_goal_expression Syntax.read_term | 
| 
029162bc9793
prefer conventional read/check distinction over manual check
 haftmann parents: 
61773diff
changeset | 112 | Syntax.read_props Attrib.check_src expression; | 
| 61669 | 113 | |
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 114 | end; | 
| 61669 | 115 | |
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 116 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 117 | (* interpretation machinery *) | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 118 | |
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 119 | local | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 120 | |
| 61672 
87203a0f0041
tuned -- share implementations as far as appropriate
 haftmann parents: 
61670diff
changeset | 121 | fun meta_rewrite eqns ctxt = | 
| 
87203a0f0041
tuned -- share implementations as far as appropriate
 haftmann parents: 
61670diff
changeset | 122 | (map (Local_Defs.meta_rewrite_rule ctxt #> Drule.abs_def) (maps snd eqns), ctxt); | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 123 | |
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 124 | fun note_eqns_register note activate deps witss def_eqns eqns attrss export export' ctxt = | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 125 | let | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 126 | val facts = (Attrib.empty_binding, [(map (Morphism.thm (export' $> export)) def_eqns, [])]) | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 127 | :: map2 (fn attrs => fn eqn => (attrs, [([Morphism.thm (export' $> export) eqn], [])])) attrss eqns; | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 128 | val (eqns', ctxt') = ctxt | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 129 | |> note Thm.theoremK facts | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 130 | |-> meta_rewrite; | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 131 | val dep_morphs = map2 (fn (dep, morph) => fn wits => | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 132 | (dep, morph $> Element.satisfy_morphism (map (Element.transform_witness export') wits))) deps witss; | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 133 | fun activate' dep_morph ctxt = activate dep_morph | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 134 | (Option.map (rpair true) (Element.eq_morphism (Proof_Context.theory_of ctxt) eqns')) export ctxt; | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 135 | in | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 136 | ctxt' | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 137 | |> fold activate' dep_morphs | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 138 | end; | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 139 | |
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 140 | in | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 141 | |
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 142 | fun generic_interpretation_with_defs prep_interpretation setup_proof note add_registration | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 143 | expression raw_defs raw_eqns initial_ctxt = | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 144 | let | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 145 | val (((propss, deps, export, export'), (def_eqns, eqns, attrss)), goal_ctxt) = | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 146 | prep_interpretation expression raw_defs raw_eqns initial_ctxt; | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 147 | fun after_qed witss eqns = | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 148 | note_eqns_register note add_registration deps witss def_eqns eqns attrss export export'; | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 149 | in setup_proof after_qed propss eqns goal_ctxt end; | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 150 | |
| 61672 
87203a0f0041
tuned -- share implementations as far as appropriate
 haftmann parents: 
61670diff
changeset | 151 | fun generic_interpretation prep_interpretation setup_proof note add_registration expression = | 
| 
87203a0f0041
tuned -- share implementations as far as appropriate
 haftmann parents: 
61670diff
changeset | 152 | generic_interpretation_with_defs prep_interpretation setup_proof note add_registration expression []; | 
| 
87203a0f0041
tuned -- share implementations as far as appropriate
 haftmann parents: 
61670diff
changeset | 153 | |
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 154 | end; | 
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 155 | |
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 156 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 157 | (** interfaces **) | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 158 | |
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 159 | (* interpretation in proofs *) | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 160 | |
| 61672 
87203a0f0041
tuned -- share implementations as far as appropriate
 haftmann parents: 
61670diff
changeset | 161 | local | 
| 
87203a0f0041
tuned -- share implementations as far as appropriate
 haftmann parents: 
61670diff
changeset | 162 | |
| 61669 | 163 | fun gen_interpret prep_interpretation expression raw_eqns int state = | 
| 164 | let | |
| 165 | val _ = Proof.assert_forward_or_chain state; | |
| 166 | val ctxt = Proof.context_of state; | |
| 167 | fun lift_after_qed after_qed witss eqns = | |
| 168 | Proof.map_context (after_qed witss eqns) #> Proof.reset_facts; | |
| 169 | fun setup_proof after_qed propss eqns goal_ctxt = | |
| 170 | Element.witness_local_proof_eqs (lift_after_qed after_qed) "interpret" | |
| 171 | propss eqns goal_ctxt int state; | |
| 172 | in | |
| 173 | generic_interpretation prep_interpretation setup_proof | |
| 174 | Attrib.local_notes (Context.proof_map ooo Locale.add_registration) expression raw_eqns ctxt | |
| 175 | end; | |
| 176 | ||
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 177 | in | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 178 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 179 | fun interpret expression = gen_interpret cert_interpretation expression; | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 180 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 181 | fun interpret_cmd raw_expression = gen_interpret read_interpretation raw_expression; | 
| 61669 | 182 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 183 | end; | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 184 | |
| 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 185 | |
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 186 | (* interpretation in local theories *) | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 187 | |
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 188 | fun interpretation expression = | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 189 | generic_interpretation cert_interpretation Element.witness_proof_eqs | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 190 | Local_Theory.notes_kind Locale.activate_fragment expression; | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 191 | |
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 192 | fun interpretation_cmd raw_expression = | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 193 | generic_interpretation read_interpretation Element.witness_proof_eqs | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 194 | Local_Theory.notes_kind Locale.activate_fragment raw_expression; | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 195 | |
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 196 | |
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 197 | (* interpretation into global theories *) | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 198 | |
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 199 | fun global_interpretation expression = | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 200 | generic_interpretation_with_defs cert_interpretation Element.witness_proof_eqs | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 201 | Local_Theory.notes_kind Local_Theory.theory_registration expression; | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 202 | |
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 203 | fun global_interpretation_cmd raw_expression = | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 204 | generic_interpretation_with_defs read_interpretation Element.witness_proof_eqs | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 205 | Local_Theory.notes_kind Local_Theory.theory_registration raw_expression; | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 206 | |
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 207 | |
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 208 | (* interpretation between locales *) | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 209 | |
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 210 | fun sublocale expression = | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 211 | generic_interpretation_with_defs cert_interpretation Element.witness_proof_eqs | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 212 | Local_Theory.notes_kind Local_Theory.locale_dependency expression; | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 213 | |
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 214 | fun sublocale_cmd raw_expression = | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 215 | generic_interpretation_with_defs read_interpretation Element.witness_proof_eqs | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 216 | Local_Theory.notes_kind Local_Theory.locale_dependency raw_expression; | 
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 217 | |
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 218 | local | 
| 61669 | 219 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 220 | fun gen_global_sublocale prep_loc prep_interpretation | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 221 | raw_locale expression raw_defs raw_eqns thy = | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 222 | let | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 223 | val lthy = Named_Target.init (prep_loc thy raw_locale) thy; | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 224 | fun setup_proof after_qed = | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 225 | Element.witness_proof_eqs | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 226 | (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 | 227 | in | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 228 | lthy |> | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 229 | generic_interpretation_with_defs prep_interpretation setup_proof | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 230 | Local_Theory.notes_kind Local_Theory.locale_dependency expression raw_defs raw_eqns | 
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 231 | end; | 
| 61669 | 232 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 233 | in | 
| 61669 | 234 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 235 | fun global_sublocale expression = | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 236 | gen_global_sublocale (K I) cert_interpretation expression; | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 237 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 238 | fun global_sublocale_cmd raw_expression = | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 239 | gen_global_sublocale Locale.check read_interpretation raw_expression; | 
| 61669 | 240 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 241 | end; | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 242 | |
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 243 | |
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 244 | (* mixed Isar interface *) | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 245 | |
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 246 | local | 
| 61670 
301e0b4ecd45
coalesce permanent_interpretation.ML with interpretation.ML
 haftmann parents: 
61669diff
changeset | 247 | |
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 248 | fun register_or_activate lthy = | 
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 249 | if Named_Target.is_theory lthy | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 250 | then Local_Theory.theory_registration | 
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 251 | else Locale.activate_fragment; | 
| 61669 | 252 | |
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 253 | fun gen_isar_interpretation prep_interpretation expression raw_eqns lthy = | 
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 254 | generic_interpretation prep_interpretation Element.witness_proof_eqs | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 255 | Local_Theory.notes_kind (register_or_activate lthy) expression raw_eqns lthy; | 
| 61669 | 256 | |
| 61673 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 257 | in | 
| 
fd4ac1530d63
represent both algebraic and local-theory views on locale interpretation in interfaces
 haftmann parents: 
61672diff
changeset | 258 | |
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 259 | fun isar_interpretation expression = | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 260 | gen_isar_interpretation cert_interpretation expression; | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 261 | fun isar_interpretation_cmd raw_expression = | 
| 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61823diff
changeset | 262 | gen_isar_interpretation read_interpretation raw_expression; | 
| 61669 | 263 | |
| 264 | end; | |
| 265 | ||
| 266 | end; |