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