| author | ballarin | 
| Sat, 19 Sep 2009 18:43:11 +0200 | |
| changeset 32801 | 6f97a67e8da8 | 
| parent 32199 | 82c4c570310a | 
| child 33389 | bb3a5fa94a91 | 
| 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 | 
| 19259 | 10 |   datatype ('typ, 'term) stmt =
 | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28079diff
changeset | 11 |     Shows of (Attrib.binding * ('term * 'term list) list) list |
 | 
| 29578 | 12 | Obtains of (binding * ((binding * 'typ option) list * 'term list)) list | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25739diff
changeset | 13 | type statement = (string, string) stmt | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25739diff
changeset | 14 | type statement_i = (typ, term) stmt | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 15 |   datatype ('typ, 'term, 'fact) ctxt =
 | 
| 29578 | 16 | Fixes of (binding * 'typ option * mixfix) list | | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 17 | Constrains of (string * 'typ) list | | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28079diff
changeset | 18 |     Assumes of (Attrib.binding * ('term * 'term list) list) list |
 | 
| 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28079diff
changeset | 19 |     Defines of (Attrib.binding * ('term * 'term list)) list |
 | 
| 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28079diff
changeset | 20 |     Notes of string * (Attrib.binding * ('fact * Attrib.src list) list) list
 | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25739diff
changeset | 21 | type context = (string, string, Facts.ref) ctxt | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25739diff
changeset | 22 | type context_i = (typ, term, thm list) ctxt | 
| 21581 | 23 |   val facts_map: (('typ, 'term, 'fact) ctxt -> ('a, 'b, 'c) ctxt) ->
 | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28079diff
changeset | 24 |    (Attrib.binding * ('fact * Attrib.src list) list) list ->
 | 
| 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28079diff
changeset | 25 |    (Attrib.binding * ('c * Attrib.src list) list) list
 | 
| 29603 | 26 |   val map_ctxt: {binding: binding -> binding, typ: 'typ -> 'a, term: 'term -> 'b,
 | 
| 27 | pattern: 'term -> 'b, fact: 'fact -> 'c, attrib: Attrib.src -> Attrib.src} -> | |
| 28 |     ('typ, 'term, 'fact) ctxt -> ('a, 'b, 'c) ctxt
 | |
| 21528 | 29 | val map_ctxt_attrib: (Attrib.src -> Attrib.src) -> | 
| 30 |     ('typ, 'term, 'fact) ctxt -> ('typ, 'term, 'fact) ctxt
 | |
| 21481 | 31 | val morph_ctxt: morphism -> context_i -> context_i | 
| 19777 | 32 | val pretty_stmt: Proof.context -> statement_i -> Pretty.T list | 
| 33 | val pretty_ctxt: Proof.context -> context_i -> Pretty.T list | |
| 34 | val pretty_statement: Proof.context -> string -> thm -> Pretty.T | |
| 35 | type witness | |
| 29578 | 36 | val prove_witness: Proof.context -> term -> tactic -> witness | 
| 37 | val witness_proof: (witness list list -> Proof.context -> Proof.context) -> | |
| 38 | term list list -> Proof.context -> Proof.state | |
| 39 | val witness_proof_eqs: (witness list list -> thm list -> Proof.context -> Proof.context) -> | |
| 40 | term list list -> term list -> Proof.context -> Proof.state | |
| 41 | val witness_local_proof: (witness list list -> Proof.state -> Proof.state) -> | |
| 42 | string -> term list list -> Proof.context -> bool -> Proof.state -> Proof.state | |
| 21481 | 43 | val morph_witness: morphism -> witness -> witness | 
| 19777 | 44 | val conclude_witness: witness -> thm | 
| 22658 
263d42253f53
Experimental interpretation code for definitions.
 ballarin parents: 
22568diff
changeset | 45 | val pretty_witness: Proof.context -> witness -> Pretty.T | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 46 | val instT_type: typ Symtab.table -> typ -> typ | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 47 | val instT_term: typ Symtab.table -> term -> term | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 48 | val instT_thm: theory -> typ Symtab.table -> thm -> thm | 
| 21481 | 49 | val instT_morphism: theory -> typ Symtab.table -> morphism | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 50 | val inst_term: typ Symtab.table * term Symtab.table -> term -> term | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 51 | val inst_thm: theory -> typ Symtab.table * term Symtab.table -> thm -> thm | 
| 21481 | 52 | val inst_morphism: theory -> typ Symtab.table * term Symtab.table -> morphism | 
| 19777 | 53 | val satisfy_thm: witness list -> thm -> thm | 
| 21481 | 54 | val satisfy_morphism: witness list -> morphism | 
| 20264 | 55 | val satisfy_facts: witness list -> | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28079diff
changeset | 56 | (Attrib.binding * (thm list * Attrib.src list) list) list -> | 
| 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28079diff
changeset | 57 | (Attrib.binding * (thm list * Attrib.src list) list) list | 
| 21581 | 58 | val generalize_facts: Proof.context -> Proof.context -> | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28079diff
changeset | 59 | (Attrib.binding * (thm list * Attrib.src list) list) list -> | 
| 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28079diff
changeset | 60 | (Attrib.binding * (thm list * Attrib.src list) list) list | 
| 29525 | 61 | val eq_morphism: theory -> thm list -> morphism | 
| 29218 | 62 | val transfer_morphism: theory -> morphism | 
| 30775 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 63 | val init: context_i -> Context.generic -> Context.generic | 
| 30777 
9960ff945c52
simplified Element.activate(_i): singleton version;
 wenzelm parents: 
30775diff
changeset | 64 | val activate_i: context_i -> Proof.context -> context_i * Proof.context | 
| 
9960ff945c52
simplified Element.activate(_i): singleton version;
 wenzelm parents: 
30775diff
changeset | 65 | 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 | 66 | end; | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 67 | |
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 68 | structure Element: ELEMENT = | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 69 | struct | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 70 | |
| 19777 | 71 | (** language elements **) | 
| 72 | ||
| 73 | (* statement *) | |
| 19259 | 74 | |
| 75 | datatype ('typ, 'term) stmt =
 | |
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28079diff
changeset | 76 |   Shows of (Attrib.binding * ('term * 'term list) list) list |
 | 
| 29578 | 77 | Obtains of (binding * ((binding * 'typ option) list * 'term list)) list; | 
| 19259 | 78 | |
| 79 | type statement = (string, string) stmt; | |
| 80 | type statement_i = (typ, term) stmt; | |
| 81 | ||
| 82 | ||
| 19777 | 83 | (* context *) | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 84 | |
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 85 | datatype ('typ, 'term, 'fact) ctxt =
 | 
| 29578 | 86 | Fixes of (binding * 'typ option * mixfix) list | | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 87 | Constrains of (string * 'typ) list | | 
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28079diff
changeset | 88 |   Assumes of (Attrib.binding * ('term * 'term list) list) list |
 | 
| 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28079diff
changeset | 89 |   Defines of (Attrib.binding * ('term * 'term list)) list |
 | 
| 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28079diff
changeset | 90 |   Notes of string * (Attrib.binding * ('fact * Attrib.src list) list) list;
 | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 91 | |
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
25739diff
changeset | 92 | type context = (string, string, Facts.ref) ctxt; | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 93 | type context_i = (typ, term, thm list) ctxt; | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 94 | |
| 21581 | 95 | fun facts_map f facts = Notes ("", facts) |> f |> (fn Notes (_, facts') => facts');
 | 
| 96 | ||
| 29603 | 97 | fun map_ctxt {binding, typ, term, pattern, fact, attrib} =
 | 
| 98 | 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 | 99 | | Constrains xs => Constrains (xs |> map (fn (x, T) => | 
| 30585 
6b2ba4666336
use Name.of_binding for basic logical entities without name space (fixes, case names etc.);
 wenzelm parents: 
30510diff
changeset | 100 | (Name.of_binding (binding (Binding.name x)), typ T))) | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 101 | | Assumes asms => Assumes (asms |> map (fn ((a, atts), propps) => | 
| 29603 | 102 | ((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 | 103 | | Defines defs => Defines (defs |> map (fn ((a, atts), (t, ps)) => | 
| 29603 | 104 | ((binding a, map attrib atts), (term t, map pattern ps)))) | 
| 21440 | 105 | | Notes (kind, facts) => Notes (kind, facts |> map (fn ((a, atts), bs) => | 
| 28965 | 106 | ((binding a, map attrib atts), bs |> map (fn (ths, btts) => (fact ths, map attrib btts))))); | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 107 | |
| 21528 | 108 | fun map_ctxt_attrib attrib = | 
| 29603 | 109 |   map_ctxt {binding = I, typ = I, term = I, pattern = I, fact = I, attrib = attrib};
 | 
| 21528 | 110 | |
| 21481 | 111 | fun morph_ctxt phi = map_ctxt | 
| 28965 | 112 |  {binding = Morphism.binding phi,
 | 
| 21481 | 113 | typ = Morphism.typ phi, | 
| 114 | term = Morphism.term phi, | |
| 29603 | 115 | pattern = Morphism.term phi, | 
| 21521 | 116 | fact = Morphism.fact phi, | 
| 21481 | 117 | attrib = Args.morph_values phi}; | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 118 | |
| 19808 | 119 | |
| 18894 | 120 | |
| 19259 | 121 | (** pretty printing **) | 
| 122 | ||
| 19267 | 123 | fun pretty_items _ _ [] = [] | 
| 124 | | pretty_items keyword sep (x :: ys) = | |
| 125 | Pretty.block [Pretty.keyword keyword, Pretty.brk 1, x] :: | |
| 126 | map (fn y => Pretty.block [Pretty.str " ", Pretty.keyword sep, Pretty.brk 1, y]) ys; | |
| 19259 | 127 | |
| 28862 | 128 | fun pretty_name_atts ctxt (b, atts) sep = | 
| 30219 | 129 | if Binding.is_empty b andalso null atts then [] | 
| 130 | else [Pretty.block (Pretty.breaks | |
| 131 | (Pretty.str (Binding.str_of b) :: Attrib.pretty_attribs ctxt atts @ [Pretty.str sep]))]; | |
| 19259 | 132 | |
| 133 | ||
| 134 | (* pretty_stmt *) | |
| 135 | ||
| 136 | fun pretty_stmt ctxt = | |
| 137 | let | |
| 24920 | 138 | val prt_typ = Pretty.quote o Syntax.pretty_typ ctxt; | 
| 139 | val prt_term = Pretty.quote o Syntax.pretty_term ctxt; | |
| 19267 | 140 | val prt_terms = separate (Pretty.keyword "and") o map prt_term; | 
| 19259 | 141 | val prt_name_atts = pretty_name_atts ctxt; | 
| 142 | ||
| 143 | fun prt_show (a, ts) = | |
| 19267 | 144 | Pretty.block (Pretty.breaks (prt_name_atts a ":" @ prt_terms (map fst ts))); | 
| 19259 | 145 | |
| 28079 
955c42c8a5e4
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
27865diff
changeset | 146 | fun prt_var (x, SOME T) = Pretty.block | 
| 30223 
24d975352879
renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
 wenzelm parents: 
30219diff
changeset | 147 | [Pretty.str (Binding.name_of x ^ " ::"), Pretty.brk 1, prt_typ T] | 
| 
24d975352879
renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
 wenzelm parents: 
30219diff
changeset | 148 | | prt_var (x, NONE) = Pretty.str (Binding.name_of x); | 
| 26721 | 149 | val prt_vars = separate (Pretty.keyword "and") o map prt_var; | 
| 19259 | 150 | |
| 19267 | 151 | fun prt_obtain (_, ([], ts)) = Pretty.block (Pretty.breaks (prt_terms ts)) | 
| 19259 | 152 | | prt_obtain (_, (xs, ts)) = Pretty.block (Pretty.breaks | 
| 19585 | 153 | (prt_vars xs @ [Pretty.keyword "where"] @ prt_terms ts)); | 
| 19259 | 154 | in | 
| 19267 | 155 | fn Shows shows => pretty_items "shows" "and" (map prt_show shows) | 
| 156 | | Obtains obtains => pretty_items "obtains" "|" (map prt_obtain obtains) | |
| 19259 | 157 | end; | 
| 158 | ||
| 18894 | 159 | |
| 19259 | 160 | (* pretty_ctxt *) | 
| 161 | ||
| 162 | fun pretty_ctxt ctxt = | |
| 163 | let | |
| 24920 | 164 | val prt_typ = Pretty.quote o Syntax.pretty_typ ctxt; | 
| 165 | val prt_term = Pretty.quote o Syntax.pretty_term ctxt; | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
31794diff
changeset | 166 | val prt_thm = Pretty.backquote o Display.pretty_thm ctxt; | 
| 19259 | 167 | val prt_name_atts = pretty_name_atts ctxt; | 
| 168 | ||
| 19267 | 169 | fun prt_mixfix NoSyn = [] | 
| 170 | | prt_mixfix mx = [Pretty.brk 2, Syntax.pretty_mixfix mx]; | |
| 171 | ||
| 30223 
24d975352879
renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
 wenzelm parents: 
30219diff
changeset | 172 | fun prt_fix (x, SOME T, mx) = Pretty.block (Pretty.str (Binding.name_of x ^ " ::") :: | 
| 28079 
955c42c8a5e4
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
27865diff
changeset | 173 | Pretty.brk 1 :: prt_typ T :: Pretty.brk 1 :: prt_mixfix mx) | 
| 30223 
24d975352879
renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
 wenzelm parents: 
30219diff
changeset | 174 | | prt_fix (x, NONE, mx) = Pretty.block (Pretty.str (Binding.name_of x) :: | 
| 28079 
955c42c8a5e4
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
27865diff
changeset | 175 | Pretty.brk 1 :: prt_mixfix mx); | 
| 28965 | 176 | fun prt_constrain (x, T) = prt_fix (Binding.name x, SOME T, NoSyn); | 
| 18894 | 177 | |
| 19259 | 178 | fun prt_asm (a, ts) = | 
| 179 | Pretty.block (Pretty.breaks (prt_name_atts a ":" @ map (prt_term o fst) ts)); | |
| 180 | fun prt_def (a, (t, _)) = | |
| 181 | Pretty.block (Pretty.breaks (prt_name_atts a ":" @ [prt_term t])); | |
| 182 | ||
| 183 | fun prt_fact (ths, []) = map prt_thm ths | |
| 184 |       | prt_fact (ths, atts) = Pretty.enclose "(" ")"
 | |
| 21032 | 185 | (Pretty.breaks (map prt_thm ths)) :: Attrib.pretty_attribs ctxt atts; | 
| 19259 | 186 | fun prt_note (a, ths) = | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19466diff
changeset | 187 | Pretty.block (Pretty.breaks (flat (prt_name_atts a "=" :: map prt_fact ths))); | 
| 19259 | 188 | in | 
| 19267 | 189 | fn Fixes fixes => pretty_items "fixes" "and" (map prt_fix fixes) | 
| 190 | | Constrains xs => pretty_items "constrains" "and" (map prt_constrain xs) | |
| 191 | | Assumes asms => pretty_items "assumes" "and" (map prt_asm asms) | |
| 192 | | Defines defs => pretty_items "defines" "and" (map prt_def defs) | |
| 21440 | 193 |      | Notes ("", facts) => pretty_items "notes" "and" (map prt_note facts)
 | 
| 194 |      | Notes (kind, facts) => pretty_items ("notes " ^ kind) "and" (map prt_note facts)
 | |
| 19259 | 195 | end; | 
| 18894 | 196 | |
| 19267 | 197 | |
| 198 | (* pretty_statement *) | |
| 199 | ||
| 200 | local | |
| 201 | ||
| 202 | fun thm_name kind th prts = | |
| 203 | let val head = | |
| 27865 
27a8ad9612a3
moved basic thm operations from structure PureThy to Thm (cf. more_thm.ML);
 wenzelm parents: 
26721diff
changeset | 204 | if Thm.has_name_hint th then | 
| 21965 
7120ef5bc378
pretty_statement: more careful handling of name_hint;
 wenzelm parents: 
21646diff
changeset | 205 | Pretty.block [Pretty.command kind, | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30280diff
changeset | 206 | Pretty.brk 1, Pretty.str (Long_Name.base_name (Thm.get_name_hint th) ^ ":")] | 
| 21965 
7120ef5bc378
pretty_statement: more careful handling of name_hint;
 wenzelm parents: 
21646diff
changeset | 207 | else Pretty.command kind | 
| 19267 | 208 | in Pretty.block (Pretty.fbreaks (head :: prts)) end; | 
| 209 | ||
| 28965 | 210 | fun fix (x, T) = (Binding.name x, SOME T); | 
| 28079 
955c42c8a5e4
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
27865diff
changeset | 211 | |
| 19267 | 212 | fun obtain prop ctxt = | 
| 213 | let | |
| 20150 | 214 | val ((xs, prop'), ctxt') = Variable.focus prop ctxt; | 
| 215 | val As = Logic.strip_imp_prems (Thm.term_of prop'); | |
| 32199 | 216 | in ((Binding.empty, (map (fix o Term.dest_Free o Thm.term_of o #2) xs, As)), ctxt') end; | 
| 19267 | 217 | |
| 218 | in | |
| 219 | ||
| 220 | fun pretty_statement ctxt kind raw_th = | |
| 221 | let | |
| 222 | val thy = ProofContext.theory_of ctxt; | |
| 20150 | 223 | val cert = Thm.cterm_of thy; | 
| 224 | ||
| 21605 | 225 | val th = MetaSimplifier.norm_hhf raw_th; | 
| 20150 | 226 | val is_elim = ObjectLogic.is_elim th; | 
| 19267 | 227 | |
| 31794 
71af1fd6a5e4
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
 wenzelm parents: 
30777diff
changeset | 228 | val ((_, [th']), ctxt') = Variable.import true [th] (Variable.set_body false ctxt); | 
| 20150 | 229 | val prop = Thm.prop_of th'; | 
| 230 | val (prems, concl) = Logic.strip_horn prop; | |
| 231 | val concl_term = ObjectLogic.drop_judgment thy concl; | |
| 19267 | 232 | |
| 20150 | 233 | val fixes = fold_aterms (fn v as Free (x, T) => | 
| 234 | if Variable.newly_fixed ctxt' ctxt x andalso not (v aconv concl_term) | |
| 26716 
8690e75e1395
print_statement: reset body mode, i.e. invent global frees (no need for revert_skolem);
 wenzelm parents: 
26628diff
changeset | 235 | then insert (op =) (x, T) else I | _ => I) prop [] |> rev; | 
| 20150 | 236 | val (assumes, cases) = take_suffix (fn prem => | 
| 237 | is_elim andalso concl aconv Logic.strip_assums_concl prem) prems; | |
| 19267 | 238 | in | 
| 28965 | 239 | pretty_ctxt ctxt' (Fixes (map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) fixes)) @ | 
| 240 | pretty_ctxt ctxt' (Assumes (map (fn t => (Attrib.empty_binding, [(t, [])])) assumes)) @ | |
| 241 | (if null cases then pretty_stmt ctxt' (Shows [(Attrib.empty_binding, [(concl, [])])]) | |
| 26716 
8690e75e1395
print_statement: reset body mode, i.e. invent global frees (no need for revert_skolem);
 wenzelm parents: 
26628diff
changeset | 242 | else | 
| 
8690e75e1395
print_statement: reset body mode, i.e. invent global frees (no need for revert_skolem);
 wenzelm parents: 
26628diff
changeset | 243 | let val (clauses, ctxt'') = fold_map (obtain o cert) cases ctxt' | 
| 
8690e75e1395
print_statement: reset body mode, i.e. invent global frees (no need for revert_skolem);
 wenzelm parents: 
26628diff
changeset | 244 | in pretty_stmt ctxt'' (Obtains clauses) end) | 
| 19267 | 245 | end |> thm_name kind raw_th; | 
| 246 | ||
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 247 | end; | 
| 19267 | 248 | |
| 19777 | 249 | |
| 250 | ||
| 251 | (** logical operations **) | |
| 252 | ||
| 253 | (* witnesses -- hypotheses as protected facts *) | |
| 254 | ||
| 255 | datatype witness = Witness of term * thm; | |
| 256 | ||
| 29578 | 257 | val mark_witness = Logic.protect; | 
| 258 | fun witness_prop (Witness (t, _)) = t; | |
| 259 | fun witness_hyps (Witness (_, th)) = #hyps (Thm.rep_thm th); | |
| 19777 | 260 | fun map_witness f (Witness witn) = Witness (f witn); | 
| 261 | ||
| 21481 | 262 | fun morph_witness phi = map_witness (fn (t, th) => (Morphism.term phi t, Morphism.thm phi th)); | 
| 263 | ||
| 20058 | 264 | fun prove_witness ctxt t tac = | 
| 29578 | 265 | Witness (t, Thm.close_derivation (Goal.prove ctxt [] [] (mark_witness t) (fn _ => | 
| 25202 
3a539d9995fb
proven witness: proper Goal.close_result save huge amounts of resources when using proof terms;
 wenzelm parents: 
24920diff
changeset | 266 | Tactic.rtac Drule.protectI 1 THEN tac))); | 
| 19777 | 267 | |
| 29603 | 268 | |
| 29578 | 269 | local | 
| 270 | ||
| 271 | val refine_witness = | |
| 30510 
4120fc59dd85
unified type Proof.method and pervasive METHOD combinators;
 wenzelm parents: 
30438diff
changeset | 272 | Proof.refine (Method.Basic (K (RAW_METHOD | 
| 29578 | 273 | (K (ALLGOALS | 
| 274 | (CONJUNCTS (ALLGOALS | |
| 32194 | 275 | (CONJUNCTS (TRYALL (Tactic.rtac Drule.protectI)))))))))); | 
| 25624 | 276 | |
| 29578 | 277 | fun gen_witness_proof proof after_qed wit_propss eq_props = | 
| 278 | let | |
| 279 | val propss = (map o map) (fn prop => (mark_witness prop, [])) wit_propss | |
| 280 | @ [map (rpair []) eq_props]; | |
| 281 | fun after_qed' thmss = | |
| 29603 | 282 | let val (wits, eqs) = split_last ((map o map) Thm.close_derivation thmss); | 
| 29578 | 283 | in after_qed ((map2 o map2) (curry Witness) wit_propss wits) eqs end; | 
| 284 | in proof after_qed' propss #> refine_witness #> Seq.hd end; | |
| 285 | ||
| 286 | in | |
| 287 | ||
| 288 | fun witness_proof after_qed wit_propss = | |
| 289 | gen_witness_proof (Proof.theorem_i NONE) (fn wits => fn _ => after_qed wits) | |
| 290 | wit_propss []; | |
| 291 | ||
| 292 | val witness_proof_eqs = gen_witness_proof (Proof.theorem_i NONE); | |
| 293 | ||
| 294 | fun witness_local_proof after_qed cmd wit_propss goal_ctxt int = | |
| 295 | gen_witness_proof (fn after_qed' => fn propss => | |
| 296 | Proof.map_context (K goal_ctxt) | |
| 297 | #> Proof.local_goal (ProofDisplay.print_results int) (K I) ProofContext.bind_propp_i | |
| 30211 | 298 | cmd NONE after_qed' (map (pair Thm.empty_binding) propss)) | 
| 29578 | 299 | (fn wits => fn _ => after_qed wits) wit_propss []; | 
| 300 | ||
| 29603 | 301 | end; | 
| 302 | ||
| 19777 | 303 | |
| 25302 | 304 | fun compose_witness (Witness (_, th)) r = | 
| 305 | let | |
| 306 | val th' = Goal.conclude th; | |
| 307 | val A = Thm.cprem_of r 1; | |
| 25739 | 308 | in | 
| 309 | Thm.implies_elim | |
| 310 | (Conv.gconv_rule Drule.beta_eta_conversion 1 r) | |
| 311 | (Conv.fconv_rule Drule.beta_eta_conversion | |
| 312 | (Thm.instantiate (Thm.match (Thm.cprop_of th', A)) th')) | |
| 313 | end; | |
| 25302 | 314 | |
| 29578 | 315 | fun conclude_witness (Witness (_, th)) = | 
| 316 | Thm.close_derivation (MetaSimplifier.norm_hhf_protect (Goal.conclude th)); | |
| 19777 | 317 | |
| 22658 
263d42253f53
Experimental interpretation code for definitions.
 ballarin parents: 
22568diff
changeset | 318 | fun pretty_witness ctxt witn = | 
| 24920 | 319 | let val prt_term = Pretty.quote o Syntax.pretty_term ctxt in | 
| 22658 
263d42253f53
Experimental interpretation code for definitions.
 ballarin parents: 
22568diff
changeset | 320 | Pretty.block (prt_term (witness_prop witn) :: | 
| 
263d42253f53
Experimental interpretation code for definitions.
 ballarin parents: 
22568diff
changeset | 321 | (if ! show_hyps then [Pretty.brk 2, Pretty.list "[" "]" | 
| 
263d42253f53
Experimental interpretation code for definitions.
 ballarin parents: 
22568diff
changeset | 322 | (map prt_term (witness_hyps witn))] else [])) | 
| 
263d42253f53
Experimental interpretation code for definitions.
 ballarin parents: 
22568diff
changeset | 323 | end; | 
| 
263d42253f53
Experimental interpretation code for definitions.
 ballarin parents: 
22568diff
changeset | 324 | |
| 19777 | 325 | |
| 326 | (* derived rules *) | |
| 327 | ||
| 20007 | 328 | fun instantiate_tfrees thy subst th = | 
| 19777 | 329 | let | 
| 330 | val certT = Thm.ctyp_of thy; | |
| 20007 | 331 | val idx = Thm.maxidx_of th + 1; | 
| 332 | fun cert_inst (a, (S, T)) = (certT (TVar ((a, idx), S)), certT T); | |
| 333 | ||
| 334 | fun add_inst (a, S) insts = | |
| 335 | if AList.defined (op =) insts a then insts | |
| 336 | else (case AList.lookup (op =) subst a of NONE => insts | SOME T => (a, (S, T)) :: insts); | |
| 337 | val insts = | |
| 338 | Term.fold_types (Term.fold_atyps (fn TFree v => add_inst v | _ => I)) | |
| 339 | (Thm.full_prop_of th) []; | |
| 19777 | 340 | in | 
| 20007 | 341 | th | 
| 342 | |> Thm.generalize (map fst insts, []) idx | |
| 343 | |> Thm.instantiate (map cert_inst insts, []) | |
| 19777 | 344 | end; | 
| 345 | ||
| 346 | fun instantiate_frees thy subst = | |
| 347 | let val cert = Thm.cterm_of thy in | |
| 348 | Drule.forall_intr_list (map (cert o Free o fst) subst) #> | |
| 349 | Drule.forall_elim_list (map (cert o snd) subst) | |
| 350 | end; | |
| 351 | ||
| 352 | fun hyps_rule rule th = | |
| 21521 | 353 |   let val {hyps, ...} = Thm.crep_thm th in
 | 
| 19777 | 354 | Drule.implies_elim_list | 
| 355 | (rule (Drule.implies_intr_list hyps th)) | |
| 21521 | 356 | (map (Thm.assume o Drule.cterm_rule rule) hyps) | 
| 19777 | 357 | end; | 
| 358 | ||
| 359 | ||
| 360 | (* instantiate types *) | |
| 361 | ||
| 362 | fun instT_type env = | |
| 363 | if Symtab.is_empty env then I | |
| 364 | else Term.map_type_tfree (fn (x, S) => the_default (TFree (x, S)) (Symtab.lookup env x)); | |
| 365 | ||
| 366 | fun instT_term env = | |
| 367 | if Symtab.is_empty env then I | |
| 20548 
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
 wenzelm parents: 
20304diff
changeset | 368 | else Term.map_types (instT_type env); | 
| 19777 | 369 | |
| 22691 | 370 | fun instT_subst env th = (Thm.fold_terms o Term.fold_types o Term.fold_atyps) | 
| 20304 | 371 | (fn T as TFree (a, _) => | 
| 372 | let val T' = the_default T (Symtab.lookup env a) | |
| 373 | in if T = T' then I else insert (op =) (a, T') end | |
| 374 | | _ => I) th []; | |
| 19777 | 375 | |
| 376 | fun instT_thm thy env th = | |
| 377 | if Symtab.is_empty env then th | |
| 378 | else | |
| 379 | let val subst = instT_subst env th | |
| 380 | in if null subst then th else th |> hyps_rule (instantiate_tfrees thy subst) end; | |
| 381 | ||
| 22672 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 382 | fun instT_morphism thy env = | 
| 24137 
8d7896398147
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 wenzelm parents: 
23414diff
changeset | 383 | let val thy_ref = Theory.check_thy thy in | 
| 22672 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 384 | Morphism.morphism | 
| 29603 | 385 |      {binding = I,
 | 
| 22672 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 386 | typ = instT_type env, | 
| 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 387 | term = instT_term env, | 
| 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 388 | fact = map (fn th => instT_thm (Theory.deref thy_ref) env th)} | 
| 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 389 | end; | 
| 19777 | 390 | |
| 391 | ||
| 392 | (* instantiate types and terms *) | |
| 393 | ||
| 394 | fun inst_term (envT, env) = | |
| 395 | if Symtab.is_empty env then instT_term envT | |
| 396 | else | |
| 397 | let | |
| 398 | val instT = instT_type envT; | |
| 399 | fun inst (Const (x, T)) = Const (x, instT T) | |
| 400 | | inst (Free (x, T)) = | |
| 401 | (case Symtab.lookup env x of | |
| 402 | NONE => Free (x, instT T) | |
| 403 | | SOME t => t) | |
| 404 | | inst (Var (xi, T)) = Var (xi, instT T) | |
| 405 | | inst (b as Bound _) = b | |
| 406 | | inst (Abs (x, T, t)) = Abs (x, instT T, inst t) | |
| 407 | | inst (t $ u) = inst t $ inst u; | |
| 408 | in Envir.beta_norm o inst end; | |
| 409 | ||
| 410 | fun inst_thm thy (envT, env) th = | |
| 411 | if Symtab.is_empty env then instT_thm thy envT th | |
| 412 | else | |
| 413 | let | |
| 414 | val substT = instT_subst envT th; | |
| 22691 | 415 | val subst = (Thm.fold_terms o Term.fold_aterms) | 
| 20304 | 416 | (fn Free (x, T) => | 
| 19777 | 417 | let | 
| 418 | val T' = instT_type envT T; | |
| 419 | val t = Free (x, T'); | |
| 420 | val t' = the_default t (Symtab.lookup env x); | |
| 20304 | 421 | in if t aconv t' then I else insert (eq_fst (op =)) ((x, T'), t') end | 
| 422 | | _ => I) th []; | |
| 19777 | 423 | in | 
| 424 | if null substT andalso null subst then th | |
| 425 | else th |> hyps_rule | |
| 426 | (instantiate_tfrees thy substT #> | |
| 427 | instantiate_frees thy subst #> | |
| 22900 | 428 | Conv.fconv_rule (Thm.beta_conversion true)) | 
| 19777 | 429 | end; | 
| 430 | ||
| 22672 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 431 | fun inst_morphism thy envs = | 
| 24137 
8d7896398147
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 wenzelm parents: 
23414diff
changeset | 432 | let val thy_ref = Theory.check_thy thy in | 
| 22672 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 433 | Morphism.morphism | 
| 29603 | 434 |      {binding = I,
 | 
| 22672 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 435 | typ = instT_type (#1 envs), | 
| 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 436 | term = inst_term envs, | 
| 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 437 | fact = map (fn th => inst_thm (Theory.deref thy_ref) envs th)} | 
| 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 438 | end; | 
| 19777 | 439 | |
| 440 | ||
| 441 | (* satisfy hypotheses *) | |
| 442 | ||
| 443 | fun satisfy_thm witns thm = thm |> fold (fn hyp => | |
| 444 | (case find_first (fn Witness (t, _) => Thm.term_of hyp aconv t) witns of | |
| 445 | NONE => I | |
| 25302 | 446 | | SOME w => Thm.implies_intr hyp #> compose_witness w)) (#hyps (Thm.crep_thm thm)); | 
| 19777 | 447 | |
| 29603 | 448 | val satisfy_morphism = Morphism.thm_morphism o satisfy_thm; | 
| 449 | val satisfy_facts = facts_map o morph_ctxt o satisfy_morphism; | |
| 20264 | 450 | |
| 451 | ||
| 29525 | 452 | (* rewriting with equalities *) | 
| 453 | ||
| 454 | fun eq_morphism thy thms = Morphism.morphism | |
| 29603 | 455 |  {binding = I,
 | 
| 456 | typ = I, | |
| 29525 | 457 | term = MetaSimplifier.rewrite_term thy thms [], | 
| 458 | fact = map (MetaSimplifier.rewrite_rule thms)}; | |
| 459 | ||
| 460 | ||
| 20264 | 461 | (* generalize type/term parameters *) | 
| 462 | ||
| 463 | val maxidx_atts = fold Args.maxidx_values; | |
| 464 | ||
| 21581 | 465 | fun generalize_facts inner outer facts = | 
| 20264 | 466 | let | 
| 467 | val thy = ProofContext.theory_of inner; | |
| 468 | val maxidx = | |
| 469 | fold (fn ((_, atts), bs) => maxidx_atts atts #> fold (maxidx_atts o #2) bs) facts ~1; | |
| 21581 | 470 | val exp_fact = map (Thm.adjust_maxidx_thm maxidx) #> Variable.export inner outer; | 
| 21521 | 471 | val exp_term = Drule.term_rule thy (singleton exp_fact); | 
| 472 | val exp_typ = Logic.type_map exp_term; | |
| 29603 | 473 |     val morphism = Morphism.morphism {binding = I, typ = exp_typ, term = exp_term, fact = exp_fact};
 | 
| 21581 | 474 | in facts_map (morph_ctxt morphism) facts end; | 
| 20886 
f26672c248ee
replaced generalize_facts by full export_(standard_)facts;
 wenzelm parents: 
20548diff
changeset | 475 | |
| 28832 | 476 | |
| 29218 | 477 | (* transfer to theory using closure *) | 
| 478 | ||
| 479 | fun transfer_morphism thy = | |
| 29603 | 480 | let val thy_ref = Theory.check_thy thy | 
| 481 | in Morphism.thm_morphism (fn th => transfer (Theory.deref thy_ref) th) end; | |
| 482 | ||
| 29218 | 483 | |
| 484 | ||
| 30775 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 485 | (** activate in context **) | 
| 28832 | 486 | |
| 30775 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 487 | (* init *) | 
| 28832 | 488 | |
| 30775 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 489 | fun init (Fixes fixes) = Context.map_proof (ProofContext.add_fixes fixes #> #2) | 
| 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 490 | | init (Constrains _) = I | 
| 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 491 | | init (Assumes asms) = Context.map_proof (fn ctxt => | 
| 28832 | 492 | let | 
| 493 | val asms' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of 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 | 494 | val (_, ctxt') = ctxt | 
| 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 495 | |> fold Variable.auto_fixes (maps (map #1 o #2) asms') | 
| 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 496 | |> ProofContext.add_assms_i Assumption.assume_export asms'; | 
| 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 497 | in ctxt' end) | 
| 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 498 | | init (Defines defs) = Context.map_proof (fn ctxt => | 
| 28832 | 499 | let | 
| 500 | val defs' = Attrib.map_specs (Attrib.attribute_i (ProofContext.theory_of ctxt)) defs; | |
| 501 | val asms = defs' |> map (fn ((name, atts), (t, ps)) => | |
| 30775 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 502 | let val ((c, _), t') = LocalDefs.cert_def ctxt t (* FIXME adapt ps? *) | 
| 30434 | 503 | in (t', ((Thm.def_binding_optional (Binding.name c) name, atts), [(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 | 504 | val (_, ctxt') = ctxt | 
| 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 505 | |> fold Variable.auto_fixes (map #1 asms) | 
| 28832 | 506 | |> ProofContext.add_assms_i LocalDefs.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 | 507 | in ctxt' end) | 
| 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 508 | | init (Notes (kind, facts)) = (fn context => | 
| 28832 | 509 | let | 
| 30775 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 510 | val facts' = Attrib.map_facts (Attrib.attribute_i (Context.theory_of context)) facts; | 
| 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 511 | val context' = context |> Context.mapping | 
| 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 512 | (PureThy.note_thmss kind facts' #> #2) | 
| 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 513 | (ProofContext.note_thmss kind facts' #> #2); | 
| 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 514 | in context' end); | 
| 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 515 | |
| 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 516 | |
| 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 517 | (* activate *) | 
| 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 518 | |
| 30777 
9960ff945c52
simplified Element.activate(_i): singleton version;
 wenzelm parents: 
30775diff
changeset | 519 | fun activate_i elem ctxt = | 
| 28832 | 520 | let | 
| 30777 
9960ff945c52
simplified Element.activate(_i): singleton version;
 wenzelm parents: 
30775diff
changeset | 521 | val elem' = map_ctxt_attrib Args.assignable elem; | 
| 
9960ff945c52
simplified Element.activate(_i): singleton version;
 wenzelm parents: 
30775diff
changeset | 522 | val ctxt' = Context.proof_map (init elem') ctxt; | 
| 
9960ff945c52
simplified Element.activate(_i): singleton version;
 wenzelm parents: 
30775diff
changeset | 523 | in (map_ctxt_attrib Args.closure elem', ctxt') end; | 
| 28832 | 524 | |
| 30777 
9960ff945c52
simplified Element.activate(_i): singleton version;
 wenzelm parents: 
30775diff
changeset | 525 | fun activate raw_elem ctxt = | 
| 
9960ff945c52
simplified Element.activate(_i): singleton version;
 wenzelm parents: 
30775diff
changeset | 526 | let val elem = raw_elem |> map_ctxt | 
| 30775 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 527 |    {binding = tap Name.of_binding,
 | 
| 29603 | 528 | typ = I, | 
| 529 | term = I, | |
| 530 | pattern = I, | |
| 30775 
71f777103225
added Element.init, which unifies former activate_elem in element.ML and init_elem in locale.ML;
 wenzelm parents: 
30763diff
changeset | 531 | fact = ProofContext.get_fact ctxt, | 
| 30777 
9960ff945c52
simplified Element.activate(_i): singleton version;
 wenzelm parents: 
30775diff
changeset | 532 | attrib = Attrib.intern_src (ProofContext.theory_of ctxt)} | 
| 
9960ff945c52
simplified Element.activate(_i): singleton version;
 wenzelm parents: 
30775diff
changeset | 533 | in activate_i elem ctxt end; | 
| 28832 | 534 | |
| 19267 | 535 | end; |