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