| author | chaieb | 
| Wed, 27 Feb 2008 14:39:49 +0100 | |
| changeset 26155 | 7c265e3da23c | 
| parent 25739 | 9da2343deb92 | 
| child 26336 | a0e2b706ce73 | 
| 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 | ID: $Id$ | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 3 | Author: Makarius | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 4 | |
| 19777 | 5 | Explicit data structures for some Isar language elements, with derived | 
| 6 | logical operations. | |
| 18140 
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 | |
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 9 | signature ELEMENT = | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 10 | sig | 
| 19259 | 11 |   datatype ('typ, 'term) stmt =
 | 
| 19585 | 12 |     Shows of ((string * Attrib.src list) * ('term * 'term list) list) list |
 | 
| 19259 | 13 | Obtains of (string * ((string * 'typ option) list * 'term list)) list | 
| 14 | type statement (*= (string, string) stmt*) | |
| 15 | type statement_i (*= (typ, term) stmt*) | |
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 16 |   datatype ('typ, 'term, 'fact) ctxt =
 | 
| 18669 | 17 | Fixes of (string * 'typ option * mixfix) list | | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 18 | Constrains of (string * 'typ) list | | 
| 19585 | 19 |     Assumes of ((string * Attrib.src list) * ('term * 'term list) list) list |
 | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 20 |     Defines of ((string * Attrib.src list) * ('term * 'term list)) list |
 | 
| 21440 | 21 |     Notes of string * ((string * Attrib.src list) * ('fact * Attrib.src list) list) list
 | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 22 | type context (*= (string, string, thmref) ctxt*) | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 23 | type context_i (*= (typ, term, thm list) ctxt*) | 
| 21581 | 24 |   val facts_map: (('typ, 'term, 'fact) ctxt -> ('a, 'b, 'c) ctxt) ->
 | 
| 25 |    ((string * Attrib.src list) * ('fact * Attrib.src list) list) list ->
 | |
| 26 |    ((string * Attrib.src list) * ('c * Attrib.src list) list) list
 | |
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 27 |   val map_ctxt: {name: string -> string,
 | 
| 18669 | 28 | var: string * mixfix -> string * mixfix, | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 29 | typ: 'typ -> 'a, term: 'term -> 'b, fact: 'fact -> 'c, | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 30 |     attrib: Attrib.src -> Attrib.src} -> ('typ, 'term, 'fact) ctxt -> ('a, 'b, 'c) ctxt
 | 
| 21528 | 31 | val map_ctxt_attrib: (Attrib.src -> Attrib.src) -> | 
| 32 |     ('typ, 'term, 'fact) ctxt -> ('typ, 'term, 'fact) ctxt
 | |
| 21481 | 33 | val morph_ctxt: morphism -> context_i -> context_i | 
| 19808 | 34 | val params_of: context_i -> (string * typ) list | 
| 35 | val prems_of: context_i -> term list | |
| 36 | val facts_of: theory -> context_i -> | |
| 37 | ((string * Attrib.src list) * (thm list * Attrib.src list) list) list | |
| 19777 | 38 | val pretty_stmt: Proof.context -> statement_i -> Pretty.T list | 
| 39 | val pretty_ctxt: Proof.context -> context_i -> Pretty.T list | |
| 40 | val pretty_statement: Proof.context -> string -> thm -> Pretty.T | |
| 41 | type witness | |
| 42 | val map_witness: (term * thm -> term * thm) -> witness -> witness | |
| 21481 | 43 | val morph_witness: morphism -> witness -> witness | 
| 19777 | 44 | val witness_prop: witness -> term | 
| 45 | val witness_hyps: witness -> term list | |
| 46 | val assume_witness: theory -> term -> witness | |
| 20058 | 47 | val prove_witness: Proof.context -> term -> tactic -> witness | 
| 25624 | 48 | val close_witness: witness -> witness | 
| 19777 | 49 | val conclude_witness: witness -> thm | 
| 50 | val mark_witness: term -> term | |
| 51 | val make_witness: term -> thm -> witness | |
| 19931 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19897diff
changeset | 52 | val dest_witness: witness -> term * thm | 
| 20068 
19c7361db4a3
New function transfer_witness lifting Thm.transfer to witnesses.
 ballarin parents: 
20058diff
changeset | 53 | val transfer_witness: theory -> witness -> witness | 
| 19808 | 54 | val refine_witness: Proof.state -> Proof.state Seq.seq | 
| 22658 
263d42253f53
Experimental interpretation code for definitions.
 ballarin parents: 
22568diff
changeset | 55 | val pretty_witness: Proof.context -> witness -> Pretty.T | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 56 | val rename: (string * (string * mixfix option)) list -> string -> string | 
| 18669 | 57 | val rename_var: (string * (string * mixfix option)) list -> string * mixfix -> string * mixfix | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 58 | val rename_term: (string * (string * mixfix option)) list -> term -> term | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 59 | val rename_thm: (string * (string * mixfix option)) list -> thm -> thm | 
| 21481 | 60 | val rename_morphism: (string * (string * mixfix option)) list -> morphism | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 61 | val instT_type: typ Symtab.table -> typ -> typ | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 62 | val instT_term: typ Symtab.table -> term -> term | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 63 | val instT_thm: theory -> typ Symtab.table -> thm -> thm | 
| 21481 | 64 | val instT_morphism: theory -> typ Symtab.table -> morphism | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 65 | val inst_term: typ Symtab.table * term Symtab.table -> term -> term | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 66 | val inst_thm: theory -> typ Symtab.table * term Symtab.table -> thm -> thm | 
| 21481 | 67 | val inst_morphism: theory -> typ Symtab.table * term Symtab.table -> morphism | 
| 19777 | 68 | val satisfy_thm: witness list -> thm -> thm | 
| 21481 | 69 | val satisfy_morphism: witness list -> morphism | 
| 20264 | 70 | val satisfy_facts: witness list -> | 
| 71 | ((string * Attrib.src list) * (thm list * Attrib.src list) list) list -> | |
| 72 | ((string * Attrib.src list) * (thm list * Attrib.src list) list) list | |
| 21581 | 73 | val generalize_facts: Proof.context -> Proof.context -> | 
| 20264 | 74 | ((string * Attrib.src list) * (thm list * Attrib.src list) list) list -> | 
| 75 | ((string * Attrib.src list) * (thm list * Attrib.src list) list) list | |
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 76 | end; | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 77 | |
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 78 | structure Element: ELEMENT = | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 79 | struct | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 80 | |
| 19777 | 81 | (** language elements **) | 
| 82 | ||
| 83 | (* statement *) | |
| 19259 | 84 | |
| 85 | datatype ('typ, 'term) stmt =
 | |
| 19585 | 86 |   Shows of ((string * Attrib.src list) * ('term * 'term list) list) list |
 | 
| 19259 | 87 | Obtains of (string * ((string * 'typ option) list * 'term list)) list; | 
| 88 | ||
| 89 | type statement = (string, string) stmt; | |
| 90 | type statement_i = (typ, term) stmt; | |
| 91 | ||
| 92 | ||
| 19777 | 93 | (* context *) | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 94 | |
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 95 | datatype ('typ, 'term, 'fact) ctxt =
 | 
| 18669 | 96 | Fixes of (string * 'typ option * mixfix) list | | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 97 | Constrains of (string * 'typ) list | | 
| 19585 | 98 |   Assumes of ((string * Attrib.src list) * ('term * 'term list) list) list |
 | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 99 |   Defines of ((string * Attrib.src list) * ('term * 'term list)) list |
 | 
| 21440 | 100 |   Notes of string * ((string * Attrib.src list) * ('fact * Attrib.src list) list) list;
 | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 101 | |
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 102 | type context = (string, string, thmref) ctxt; | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 103 | type context_i = (typ, term, thm list) ctxt; | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 104 | |
| 21581 | 105 | fun facts_map f facts = Notes ("", facts) |> f |> (fn Notes (_, facts') => facts');
 | 
| 106 | ||
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 107 | fun map_ctxt {name, var, typ, term, fact, attrib} =
 | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 108 | fn Fixes fixes => Fixes (fixes |> map (fn (x, T, mx) => | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 109 | let val (x', mx') = var (x, mx) in (x', Option.map typ T, mx') end)) | 
| 18669 | 110 | | Constrains xs => Constrains (xs |> map (fn (x, T) => (#1 (var (x, NoSyn)), typ T))) | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 111 | | Assumes asms => Assumes (asms |> map (fn ((a, atts), propps) => | 
| 19585 | 112 | ((name a, map attrib atts), propps |> map (fn (t, ps) => (term t, map term ps))))) | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 113 | | Defines defs => Defines (defs |> map (fn ((a, atts), (t, ps)) => | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 114 | ((name a, map attrib atts), (term t, map term ps)))) | 
| 21440 | 115 | | Notes (kind, facts) => Notes (kind, facts |> map (fn ((a, atts), bs) => | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 116 | ((name a, map attrib atts), bs |> map (fn (ths, btts) => (fact ths, map attrib btts))))); | 
| 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 117 | |
| 21528 | 118 | fun map_ctxt_attrib attrib = | 
| 119 |   map_ctxt {name = I, var = I, typ = I, term = I, fact = I, attrib = attrib};
 | |
| 120 | ||
| 21481 | 121 | fun morph_ctxt phi = map_ctxt | 
| 122 |  {name = Morphism.name phi,
 | |
| 123 | var = Morphism.var phi, | |
| 124 | typ = Morphism.typ phi, | |
| 125 | term = Morphism.term phi, | |
| 21521 | 126 | fact = Morphism.fact phi, | 
| 21481 | 127 | attrib = Args.morph_values phi}; | 
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 128 | |
| 19808 | 129 | |
| 130 | (* logical content *) | |
| 131 | ||
| 19777 | 132 | fun params_of (Fixes fixes) = fixes |> map | 
| 133 | (fn (x, SOME T, _) => (x, T) | |
| 134 |       | (x, _, _) => raise TERM ("Untyped context element parameter " ^ quote x, []))
 | |
| 135 | | params_of _ = []; | |
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 136 | |
| 19777 | 137 | fun prems_of (Assumes asms) = maps (map fst o snd) asms | 
| 138 | | prems_of (Defines defs) = map (fst o snd) defs | |
| 139 | | prems_of _ = []; | |
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 140 | |
| 20233 | 141 | fun assume thy t = Assumption.assume (Thm.cterm_of thy t); | 
| 19808 | 142 | |
| 143 | fun facts_of thy (Assumes asms) = map (apsnd (map (fn (t, _) => ([assume thy t], [])))) asms | |
| 144 | | facts_of thy (Defines defs) = map (apsnd (fn (t, _) => [([assume thy t], [])])) defs | |
| 21440 | 145 | | facts_of _ (Notes (_, facts)) = facts | 
| 19808 | 146 | | facts_of _ _ = []; | 
| 147 | ||
| 18894 | 148 | |
| 149 | ||
| 19259 | 150 | (** pretty printing **) | 
| 151 | ||
| 19267 | 152 | fun pretty_items _ _ [] = [] | 
| 153 | | pretty_items keyword sep (x :: ys) = | |
| 154 | Pretty.block [Pretty.keyword keyword, Pretty.brk 1, x] :: | |
| 155 | map (fn y => Pretty.block [Pretty.str " ", Pretty.keyword sep, Pretty.brk 1, y]) ys; | |
| 19259 | 156 | |
| 157 | fun pretty_name_atts ctxt (name, atts) sep = | |
| 158 | if name = "" andalso null atts then [] | |
| 19731 | 159 | else [Pretty.block | 
| 21032 | 160 | (Pretty.breaks (Pretty.str name :: Attrib.pretty_attribs ctxt atts @ [Pretty.str sep]))]; | 
| 19259 | 161 | |
| 162 | ||
| 163 | (* pretty_stmt *) | |
| 164 | ||
| 165 | fun pretty_stmt ctxt = | |
| 166 | let | |
| 24920 | 167 | val prt_typ = Pretty.quote o Syntax.pretty_typ ctxt; | 
| 168 | val prt_term = Pretty.quote o Syntax.pretty_term ctxt; | |
| 19267 | 169 | val prt_terms = separate (Pretty.keyword "and") o map prt_term; | 
| 19259 | 170 | val prt_name_atts = pretty_name_atts ctxt; | 
| 171 | ||
| 172 | fun prt_show (a, ts) = | |
| 19267 | 173 | Pretty.block (Pretty.breaks (prt_name_atts a ":" @ prt_terms (map fst ts))); | 
| 19259 | 174 | |
| 175 | fun prt_var (x, SOME T) = Pretty.block [Pretty.str (x ^ " ::"), Pretty.brk 1, prt_typ T] | |
| 176 | | prt_var (x, NONE) = Pretty.str x; | |
| 19585 | 177 | val prt_vars = separate (Pretty.keyword "and") o map prt_var; | 
| 19259 | 178 | |
| 19267 | 179 | fun prt_obtain (_, ([], ts)) = Pretty.block (Pretty.breaks (prt_terms ts)) | 
| 19259 | 180 | | prt_obtain (_, (xs, ts)) = Pretty.block (Pretty.breaks | 
| 19585 | 181 | (prt_vars xs @ [Pretty.keyword "where"] @ prt_terms ts)); | 
| 19259 | 182 | in | 
| 19267 | 183 | fn Shows shows => pretty_items "shows" "and" (map prt_show shows) | 
| 184 | | Obtains obtains => pretty_items "obtains" "|" (map prt_obtain obtains) | |
| 19259 | 185 | end; | 
| 186 | ||
| 18894 | 187 | |
| 19259 | 188 | (* pretty_ctxt *) | 
| 189 | ||
| 190 | fun pretty_ctxt ctxt = | |
| 191 | let | |
| 24920 | 192 | val prt_typ = Pretty.quote o Syntax.pretty_typ ctxt; | 
| 193 | val prt_term = Pretty.quote o Syntax.pretty_term ctxt; | |
| 19259 | 194 | val prt_thm = Pretty.backquote o ProofContext.pretty_thm ctxt; | 
| 195 | val prt_name_atts = pretty_name_atts ctxt; | |
| 196 | ||
| 19267 | 197 | fun prt_mixfix NoSyn = [] | 
| 198 | | prt_mixfix mx = [Pretty.brk 2, Syntax.pretty_mixfix mx]; | |
| 199 | ||
| 19259 | 200 | fun prt_fix (x, SOME T, mx) = Pretty.block (Pretty.str (x ^ " ::") :: Pretty.brk 1 :: | 
| 201 | prt_typ T :: Pretty.brk 1 :: prt_mixfix mx) | |
| 202 | | prt_fix (x, NONE, mx) = Pretty.block (Pretty.str x :: Pretty.brk 1 :: prt_mixfix mx); | |
| 203 | fun prt_constrain (x, T) = prt_fix (x, SOME T, NoSyn); | |
| 18894 | 204 | |
| 19259 | 205 | fun prt_asm (a, ts) = | 
| 206 | Pretty.block (Pretty.breaks (prt_name_atts a ":" @ map (prt_term o fst) ts)); | |
| 207 | fun prt_def (a, (t, _)) = | |
| 208 | Pretty.block (Pretty.breaks (prt_name_atts a ":" @ [prt_term t])); | |
| 209 | ||
| 210 | fun prt_fact (ths, []) = map prt_thm ths | |
| 211 |       | prt_fact (ths, atts) = Pretty.enclose "(" ")"
 | |
| 21032 | 212 | (Pretty.breaks (map prt_thm ths)) :: Attrib.pretty_attribs ctxt atts; | 
| 19259 | 213 | fun prt_note (a, ths) = | 
| 19482 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 wenzelm parents: 
19466diff
changeset | 214 | Pretty.block (Pretty.breaks (flat (prt_name_atts a "=" :: map prt_fact ths))); | 
| 19259 | 215 | in | 
| 19267 | 216 | fn Fixes fixes => pretty_items "fixes" "and" (map prt_fix fixes) | 
| 217 | | Constrains xs => pretty_items "constrains" "and" (map prt_constrain xs) | |
| 218 | | Assumes asms => pretty_items "assumes" "and" (map prt_asm asms) | |
| 219 | | Defines defs => pretty_items "defines" "and" (map prt_def defs) | |
| 21440 | 220 |      | Notes ("", facts) => pretty_items "notes" "and" (map prt_note facts)
 | 
| 221 |      | Notes (kind, facts) => pretty_items ("notes " ^ kind) "and" (map prt_note facts)
 | |
| 19259 | 222 | end; | 
| 18894 | 223 | |
| 19267 | 224 | |
| 225 | (* pretty_statement *) | |
| 226 | ||
| 227 | local | |
| 228 | ||
| 229 | fun thm_name kind th prts = | |
| 230 | let val head = | |
| 21965 
7120ef5bc378
pretty_statement: more careful handling of name_hint;
 wenzelm parents: 
21646diff
changeset | 231 | if PureThy.has_name_hint th then | 
| 
7120ef5bc378
pretty_statement: more careful handling of name_hint;
 wenzelm parents: 
21646diff
changeset | 232 | Pretty.block [Pretty.command kind, | 
| 
7120ef5bc378
pretty_statement: more careful handling of name_hint;
 wenzelm parents: 
21646diff
changeset | 233 | Pretty.brk 1, Pretty.str (Sign.base_name (PureThy.get_name_hint th) ^ ":")] | 
| 
7120ef5bc378
pretty_statement: more careful handling of name_hint;
 wenzelm parents: 
21646diff
changeset | 234 | else Pretty.command kind | 
| 19267 | 235 | in Pretty.block (Pretty.fbreaks (head :: prts)) end; | 
| 236 | ||
| 237 | fun obtain prop ctxt = | |
| 238 | let | |
| 20150 | 239 | val ((xs, prop'), ctxt') = Variable.focus prop ctxt; | 
| 240 | val As = Logic.strip_imp_prems (Thm.term_of prop'); | |
| 241 | fun var (x, T) = (ProofContext.revert_skolem ctxt' x, SOME T); | |
| 20218 
be3bfb0699ba
Variable.import(T): result includes fixed types/terms;
 wenzelm parents: 
20150diff
changeset | 242 |   in (("", (map (var o Term.dest_Free o Thm.term_of) xs, As)), ctxt') end;
 | 
| 19267 | 243 | |
| 244 | in | |
| 245 | ||
| 246 | fun pretty_statement ctxt kind raw_th = | |
| 247 | let | |
| 248 | val thy = ProofContext.theory_of ctxt; | |
| 20150 | 249 | val cert = Thm.cterm_of thy; | 
| 250 | ||
| 21605 | 251 | val th = MetaSimplifier.norm_hhf raw_th; | 
| 20150 | 252 | val is_elim = ObjectLogic.is_elim th; | 
| 19267 | 253 | |
| 22568 
ed7aa5a350ef
renamed Variable.import to import_thms (avoid clash with Alice keywords);
 wenzelm parents: 
21965diff
changeset | 254 | val ((_, [th']), ctxt') = Variable.import_thms true [th] ctxt; | 
| 20150 | 255 | val prop = Thm.prop_of th'; | 
| 256 | val (prems, concl) = Logic.strip_horn prop; | |
| 257 | val concl_term = ObjectLogic.drop_judgment thy concl; | |
| 19267 | 258 | |
| 20150 | 259 | val fixes = fold_aterms (fn v as Free (x, T) => | 
| 260 | if Variable.newly_fixed ctxt' ctxt x andalso not (v aconv concl_term) | |
| 261 | then insert (op =) (x, T) else I | _ => I) prop [] | |
| 262 | |> rev |> map (apfst (ProofContext.revert_skolem ctxt')); | |
| 263 | val (assumes, cases) = take_suffix (fn prem => | |
| 264 | is_elim andalso concl aconv Logic.strip_assums_concl prem) prems; | |
| 19267 | 265 | in | 
| 20150 | 266 | pretty_ctxt ctxt' (Fixes (map (fn (x, T) => (x, SOME T, NoSyn)) fixes)) @ | 
| 267 |     pretty_ctxt ctxt' (Assumes (map (fn t => (("", []), [(t, [])])) assumes)) @
 | |
| 268 | pretty_stmt ctxt' | |
| 19585 | 269 |      (if null cases then Shows [(("", []), [(concl, [])])]
 | 
| 20150 | 270 | else Obtains (#1 (fold_map (obtain o cert) cases ctxt'))) | 
| 19267 | 271 | end |> thm_name kind raw_th; | 
| 272 | ||
| 18140 
691c64d615a5
Explicit data structures for some Isar language elements.
 wenzelm parents: diff
changeset | 273 | end; | 
| 19267 | 274 | |
| 19777 | 275 | |
| 276 | ||
| 277 | (** logical operations **) | |
| 278 | ||
| 279 | (* witnesses -- hypotheses as protected facts *) | |
| 280 | ||
| 281 | datatype witness = Witness of term * thm; | |
| 282 | ||
| 283 | fun map_witness f (Witness witn) = Witness (f witn); | |
| 284 | ||
| 21481 | 285 | fun morph_witness phi = map_witness (fn (t, th) => (Morphism.term phi t, Morphism.thm phi th)); | 
| 286 | ||
| 19777 | 287 | fun witness_prop (Witness (t, _)) = t; | 
| 288 | fun witness_hyps (Witness (_, th)) = #hyps (Thm.rep_thm th); | |
| 289 | ||
| 290 | fun assume_witness thy t = | |
| 291 | Witness (t, Goal.protect (Thm.assume (Thm.cterm_of thy t))); | |
| 292 | ||
| 20058 | 293 | fun prove_witness ctxt t tac = | 
| 25202 
3a539d9995fb
proven witness: proper Goal.close_result save huge amounts of resources when using proof terms;
 wenzelm parents: 
24920diff
changeset | 294 | Witness (t, Goal.close_result (Goal.prove ctxt [] [] (Logic.protect t) (fn _ => | 
| 
3a539d9995fb
proven witness: proper Goal.close_result save huge amounts of resources when using proof terms;
 wenzelm parents: 
24920diff
changeset | 295 | Tactic.rtac Drule.protectI 1 THEN tac))); | 
| 19777 | 296 | |
| 25624 | 297 | val close_witness = map_witness (fn (t, th) => (t, Goal.close_result th)); | 
| 298 | ||
| 25202 
3a539d9995fb
proven witness: proper Goal.close_result save huge amounts of resources when using proof terms;
 wenzelm parents: 
24920diff
changeset | 299 | fun conclude_witness (Witness (_, th)) = | 
| 
3a539d9995fb
proven witness: proper Goal.close_result save huge amounts of resources when using proof terms;
 wenzelm parents: 
24920diff
changeset | 300 | Goal.close_result (MetaSimplifier.norm_hhf_protect (Goal.conclude th)); | 
| 19777 | 301 | |
| 25302 | 302 | fun compose_witness (Witness (_, th)) r = | 
| 303 | let | |
| 304 | val th' = Goal.conclude th; | |
| 305 | val A = Thm.cprem_of r 1; | |
| 25739 | 306 | in | 
| 307 | Thm.implies_elim | |
| 308 | (Conv.gconv_rule Drule.beta_eta_conversion 1 r) | |
| 309 | (Conv.fconv_rule Drule.beta_eta_conversion | |
| 310 | (Thm.instantiate (Thm.match (Thm.cprop_of th', A)) th')) | |
| 311 | end; | |
| 25302 | 312 | |
| 19777 | 313 | val mark_witness = Logic.protect; | 
| 314 | ||
| 315 | fun make_witness t th = Witness (t, th); | |
| 19931 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19897diff
changeset | 316 | fun dest_witness (Witness w) = w; | 
| 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 ballarin parents: 
19897diff
changeset | 317 | |
| 20068 
19c7361db4a3
New function transfer_witness lifting Thm.transfer to witnesses.
 ballarin parents: 
20058diff
changeset | 318 | fun transfer_witness thy (Witness (t, th)) = Witness (t, Thm.transfer thy th); | 
| 
19c7361db4a3
New function transfer_witness lifting Thm.transfer to witnesses.
 ballarin parents: 
20058diff
changeset | 319 | |
| 19777 | 320 | val refine_witness = | 
| 321 | Proof.refine (Method.Basic (K (Method.RAW_METHOD | |
| 322 | (K (ALLGOALS | |
| 23414 
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
 wenzelm parents: 
23351diff
changeset | 323 | (CONJUNCTS (ALLGOALS | 
| 
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
 wenzelm parents: 
23351diff
changeset | 324 | (CONJUNCTS (TRYALL (Tactic.rtac Drule.protectI)))))))), Position.none)); | 
| 19777 | 325 | |
| 22658 
263d42253f53
Experimental interpretation code for definitions.
 ballarin parents: 
22568diff
changeset | 326 | fun pretty_witness ctxt witn = | 
| 24920 | 327 | let val prt_term = Pretty.quote o Syntax.pretty_term ctxt in | 
| 22658 
263d42253f53
Experimental interpretation code for definitions.
 ballarin parents: 
22568diff
changeset | 328 | Pretty.block (prt_term (witness_prop witn) :: | 
| 
263d42253f53
Experimental interpretation code for definitions.
 ballarin parents: 
22568diff
changeset | 329 | (if ! show_hyps then [Pretty.brk 2, Pretty.list "[" "]" | 
| 
263d42253f53
Experimental interpretation code for definitions.
 ballarin parents: 
22568diff
changeset | 330 | (map prt_term (witness_hyps witn))] else [])) | 
| 
263d42253f53
Experimental interpretation code for definitions.
 ballarin parents: 
22568diff
changeset | 331 | end; | 
| 
263d42253f53
Experimental interpretation code for definitions.
 ballarin parents: 
22568diff
changeset | 332 | |
| 19777 | 333 | |
| 334 | (* derived rules *) | |
| 335 | ||
| 20007 | 336 | fun instantiate_tfrees thy subst th = | 
| 19777 | 337 | let | 
| 338 | val certT = Thm.ctyp_of thy; | |
| 20007 | 339 | val idx = Thm.maxidx_of th + 1; | 
| 340 | fun cert_inst (a, (S, T)) = (certT (TVar ((a, idx), S)), certT T); | |
| 341 | ||
| 342 | fun add_inst (a, S) insts = | |
| 343 | if AList.defined (op =) insts a then insts | |
| 344 | else (case AList.lookup (op =) subst a of NONE => insts | SOME T => (a, (S, T)) :: insts); | |
| 345 | val insts = | |
| 346 | Term.fold_types (Term.fold_atyps (fn TFree v => add_inst v | _ => I)) | |
| 347 | (Thm.full_prop_of th) []; | |
| 19777 | 348 | in | 
| 20007 | 349 | th | 
| 350 | |> Thm.generalize (map fst insts, []) idx | |
| 351 | |> Thm.instantiate (map cert_inst insts, []) | |
| 19777 | 352 | end; | 
| 353 | ||
| 354 | fun instantiate_frees thy subst = | |
| 355 | let val cert = Thm.cterm_of thy in | |
| 356 | Drule.forall_intr_list (map (cert o Free o fst) subst) #> | |
| 357 | Drule.forall_elim_list (map (cert o snd) subst) | |
| 358 | end; | |
| 359 | ||
| 360 | fun hyps_rule rule th = | |
| 21521 | 361 |   let val {hyps, ...} = Thm.crep_thm th in
 | 
| 19777 | 362 | Drule.implies_elim_list | 
| 363 | (rule (Drule.implies_intr_list hyps th)) | |
| 21521 | 364 | (map (Thm.assume o Drule.cterm_rule rule) hyps) | 
| 19777 | 365 | end; | 
| 366 | ||
| 367 | ||
| 368 | (* rename *) | |
| 369 | ||
| 370 | fun rename ren x = | |
| 371 | (case AList.lookup (op =) ren (x: string) of | |
| 372 | NONE => x | |
| 373 | | SOME (x', _) => x'); | |
| 374 | ||
| 375 | fun rename_var ren (x, mx) = | |
| 376 | (case (AList.lookup (op =) ren (x: string), mx) of | |
| 377 | (NONE, _) => (x, mx) | |
| 378 | | (SOME (x', NONE), Structure) => (x', mx) | |
| 379 | | (SOME (x', SOME _), Structure) => | |
| 380 |       error ("Attempt to change syntax of structure parameter " ^ quote x)
 | |
| 381 | | (SOME (x', NONE), _) => (x', NoSyn) | |
| 382 | | (SOME (x', SOME mx'), _) => (x', mx')); | |
| 383 | ||
| 384 | fun rename_term ren (Free (x, T)) = Free (rename ren x, T) | |
| 385 | | rename_term ren (t $ u) = rename_term ren t $ rename_term ren u | |
| 386 | | rename_term ren (Abs (x, T, t)) = Abs (x, T, rename_term ren t) | |
| 387 | | rename_term _ a = a; | |
| 388 | ||
| 389 | fun rename_thm ren th = | |
| 390 | let | |
| 20304 | 391 | val thy = Thm.theory_of_thm th; | 
| 22691 | 392 | val subst = (Thm.fold_terms o Term.fold_aterms) | 
| 20304 | 393 | (fn Free (x, T) => | 
| 19777 | 394 | let val x' = rename ren x | 
| 20304 | 395 | in if x = x' then I else insert (eq_fst (op =)) ((x, T), Free (x', T)) end | 
| 396 | | _ => I) th []; | |
| 19777 | 397 | in | 
| 398 | if null subst then th | |
| 20304 | 399 | else th |> hyps_rule (instantiate_frees thy subst) | 
| 19777 | 400 | end; | 
| 401 | ||
| 21481 | 402 | fun rename_morphism ren = Morphism.morphism | 
| 21521 | 403 |   {name = I, var = rename_var ren, typ = I, term = rename_term ren, fact = map (rename_thm ren)};
 | 
| 19777 | 404 | |
| 405 | ||
| 406 | (* instantiate types *) | |
| 407 | ||
| 408 | fun instT_type env = | |
| 409 | if Symtab.is_empty env then I | |
| 410 | else Term.map_type_tfree (fn (x, S) => the_default (TFree (x, S)) (Symtab.lookup env x)); | |
| 411 | ||
| 412 | fun instT_term env = | |
| 413 | 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 | 414 | else Term.map_types (instT_type env); | 
| 19777 | 415 | |
| 22691 | 416 | fun instT_subst env th = (Thm.fold_terms o Term.fold_types o Term.fold_atyps) | 
| 20304 | 417 | (fn T as TFree (a, _) => | 
| 418 | let val T' = the_default T (Symtab.lookup env a) | |
| 419 | in if T = T' then I else insert (op =) (a, T') end | |
| 420 | | _ => I) th []; | |
| 19777 | 421 | |
| 422 | fun instT_thm thy env th = | |
| 423 | if Symtab.is_empty env then th | |
| 424 | else | |
| 425 | let val subst = instT_subst env th | |
| 426 | in if null subst then th else th |> hyps_rule (instantiate_tfrees thy subst) end; | |
| 427 | ||
| 22672 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 428 | 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 | 429 | let val thy_ref = Theory.check_thy thy in | 
| 22672 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 430 | Morphism.morphism | 
| 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 431 |      {name = I, var = I,
 | 
| 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 432 | typ = instT_type env, | 
| 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 433 | term = instT_term env, | 
| 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 434 | 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 | 435 | end; | 
| 19777 | 436 | |
| 437 | ||
| 438 | (* instantiate types and terms *) | |
| 439 | ||
| 440 | fun inst_term (envT, env) = | |
| 441 | if Symtab.is_empty env then instT_term envT | |
| 442 | else | |
| 443 | let | |
| 444 | val instT = instT_type envT; | |
| 445 | fun inst (Const (x, T)) = Const (x, instT T) | |
| 446 | | inst (Free (x, T)) = | |
| 447 | (case Symtab.lookup env x of | |
| 448 | NONE => Free (x, instT T) | |
| 449 | | SOME t => t) | |
| 450 | | inst (Var (xi, T)) = Var (xi, instT T) | |
| 451 | | inst (b as Bound _) = b | |
| 452 | | inst (Abs (x, T, t)) = Abs (x, instT T, inst t) | |
| 453 | | inst (t $ u) = inst t $ inst u; | |
| 454 | in Envir.beta_norm o inst end; | |
| 455 | ||
| 456 | fun inst_thm thy (envT, env) th = | |
| 457 | if Symtab.is_empty env then instT_thm thy envT th | |
| 458 | else | |
| 459 | let | |
| 460 | val substT = instT_subst envT th; | |
| 22691 | 461 | val subst = (Thm.fold_terms o Term.fold_aterms) | 
| 20304 | 462 | (fn Free (x, T) => | 
| 19777 | 463 | let | 
| 464 | val T' = instT_type envT T; | |
| 465 | val t = Free (x, T'); | |
| 466 | val t' = the_default t (Symtab.lookup env x); | |
| 20304 | 467 | in if t aconv t' then I else insert (eq_fst (op =)) ((x, T'), t') end | 
| 468 | | _ => I) th []; | |
| 19777 | 469 | in | 
| 470 | if null substT andalso null subst then th | |
| 471 | else th |> hyps_rule | |
| 472 | (instantiate_tfrees thy substT #> | |
| 473 | instantiate_frees thy subst #> | |
| 22900 | 474 | Conv.fconv_rule (Thm.beta_conversion true)) | 
| 19777 | 475 | end; | 
| 476 | ||
| 22672 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 477 | 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 | 478 | let val thy_ref = Theory.check_thy thy in | 
| 22672 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 479 | Morphism.morphism | 
| 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 480 |      {name = I, var = I,
 | 
| 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 481 | typ = instT_type (#1 envs), | 
| 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 482 | term = inst_term envs, | 
| 
777af26d5713
inst(T)_morphism: avoid reference to static theory value;
 wenzelm parents: 
22658diff
changeset | 483 | 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 | 484 | end; | 
| 19777 | 485 | |
| 486 | ||
| 487 | (* satisfy hypotheses *) | |
| 488 | ||
| 489 | fun satisfy_thm witns thm = thm |> fold (fn hyp => | |
| 490 | (case find_first (fn Witness (t, _) => Thm.term_of hyp aconv t) witns of | |
| 491 | NONE => I | |
| 25302 | 492 | | SOME w => Thm.implies_intr hyp #> compose_witness w)) (#hyps (Thm.crep_thm thm)); | 
| 19777 | 493 | |
| 21497 | 494 | fun satisfy_morphism witns = Morphism.thm_morphism (satisfy_thm witns); | 
| 19843 | 495 | |
| 21581 | 496 | fun satisfy_facts witns = facts_map (morph_ctxt (satisfy_morphism witns)); | 
| 20264 | 497 | |
| 498 | ||
| 499 | (* generalize type/term parameters *) | |
| 500 | ||
| 501 | val maxidx_atts = fold Args.maxidx_values; | |
| 502 | ||
| 21581 | 503 | fun generalize_facts inner outer facts = | 
| 20264 | 504 | let | 
| 505 | val thy = ProofContext.theory_of inner; | |
| 506 | val maxidx = | |
| 507 | fold (fn ((_, atts), bs) => maxidx_atts atts #> fold (maxidx_atts o #2) bs) facts ~1; | |
| 21581 | 508 | val exp_fact = map (Thm.adjust_maxidx_thm maxidx) #> Variable.export inner outer; | 
| 21521 | 509 | val exp_term = Drule.term_rule thy (singleton exp_fact); | 
| 510 | val exp_typ = Logic.type_map exp_term; | |
| 511 | val morphism = | |
| 512 |       Morphism.morphism {name = I, var = I, typ = exp_typ, term = exp_term, fact = exp_fact};
 | |
| 21581 | 513 | in facts_map (morph_ctxt morphism) facts end; | 
| 20886 
f26672c248ee
replaced generalize_facts by full export_(standard_)facts;
 wenzelm parents: 
20548diff
changeset | 514 | |
| 19267 | 515 | end; |