| author | wenzelm | 
| Sun, 15 Jan 2023 20:00:22 +0100 | |
| changeset 76988 | 7f7d5c93e36b | 
| parent 74561 | 8e6c973003c8 | 
| child 82967 | 73af47bc277c | 
| permissions | -rw-r--r-- | 
| 58816 | 1 | (* Title: HOL/Tools/Function/function_context_tree.ML | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 2 | Author: Alexander Krauss, TU Muenchen | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 3 | |
| 41114 | 4 | Construction and traversal of trees of nested contexts along a term. | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 5 | *) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 6 | |
| 58816 | 7 | signature FUNCTION_CONTEXT_TREE = | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 8 | sig | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 9 | (* poor man's contexts: fixes + assumes *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 10 | type ctxt = (string * typ) list * thm list | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 11 | type ctx_tree | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 12 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 13 | val get_function_congs : Proof.context -> thm list | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 14 | val add_function_cong : thm -> Context.generic -> Context.generic | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 15 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 16 | val cong_add: attribute | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 17 | val cong_del: attribute | 
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 18 | |
| 65418 
c821f1f3d92d
more general signature; works for all terms, not just frees
 Lars Hupel <lars.hupel@mytum.de> parents: 
61112diff
changeset | 19 | val mk_tree: term -> term -> Proof.context -> term -> ctx_tree | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 20 | |
| 59618 | 21 | val inst_tree: Proof.context -> term -> term -> ctx_tree -> ctx_tree | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 22 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 23 | val export_term : ctxt -> term -> term | 
| 59618 | 24 | val export_thm : Proof.context -> ctxt -> thm -> thm | 
| 25 | val import_thm : Proof.context -> ctxt -> thm -> thm | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 26 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 27 | val traverse_tree : | 
| 26115 | 28 | (ctxt -> term -> | 
| 29 | (ctxt * thm) list -> | |
| 30 | (ctxt * thm) list * 'b -> | |
| 31 | (ctxt * thm) list * 'b) | |
| 23819 | 32 | -> ctx_tree -> 'b -> 'b | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 33 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
44338diff
changeset | 34 | val rewrite_by_tree : Proof.context -> term -> thm -> (thm * thm) list -> | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 35 | ctx_tree -> thm * (thm * thm) list | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 36 | end | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 37 | |
| 58816 | 38 | structure Function_Context_Tree : FUNCTION_CONTEXT_TREE = | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 39 | struct | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 40 | |
| 26115 | 41 | type ctxt = (string * typ) list * thm list | 
| 42 | ||
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33049diff
changeset | 43 | open Function_Common | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33049diff
changeset | 44 | open Function_Lib | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 45 | |
| 33519 | 46 | structure FunctionCongs = Generic_Data | 
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 47 | ( | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 48 | type T = thm list | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 49 | val empty = [] | 
| 33519 | 50 | val merge = Thm.merge_thms | 
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 51 | ); | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 52 | |
| 61112 | 53 | fun get_function_congs ctxt = | 
| 54 | FunctionCongs.get (Context.Proof ctxt) | |
| 67649 | 55 | |> map (Thm.transfer' ctxt); | 
| 61112 | 56 | |
| 57 | val add_function_cong = FunctionCongs.map o Thm.add_thm o Thm.trim_context; | |
| 58 | ||
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 59 | |
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 60 | (* congruence rules *) | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 61 | |
| 61112 | 62 | val cong_add = Thm.declaration_attribute (add_function_cong o safe_mk_meta_eq); | 
| 63 | val cong_del = Thm.declaration_attribute (FunctionCongs.map o Thm.del_thm o safe_mk_meta_eq); | |
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 64 | |
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 65 | |
| 35403 | 66 | type depgraph = int Int_Graph.T | 
| 23819 | 67 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 68 | datatype ctx_tree = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 69 | Leaf of term | 
| 26115 | 70 | | Cong of (thm * depgraph * (ctxt * ctx_tree) list) | 
| 23819 | 71 | | RCall of (term * ctx_tree) | 
| 72 | ||
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 73 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 74 | (* Maps "Trueprop A = B" to "A" *) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 75 | val rhs_of = snd o HOLogic.dest_eq o HOLogic.dest_Trueprop | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 76 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 77 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 78 | (*** Dependency analysis for congruence rules ***) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 79 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 80 | fun branch_vars t = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 81 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 82 | val t' = snd (dest_all_all t) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 83 | val (assumes, concl) = Logic.strip_horn t' | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 84 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 85 | (fold Term.add_vars assumes [], Term.add_vars concl []) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 86 | end | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 87 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 88 | fun cong_deps crule = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 89 | let | 
| 59582 | 90 | val num_branches = map_index (apsnd branch_vars) (Thm.prems_of crule) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 91 | in | 
| 35403 | 92 | Int_Graph.empty | 
| 93 | |> fold (fn (i,_)=> Int_Graph.new_node (i,i)) num_branches | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 94 | |> fold_product (fn (i, (c1, _)) => fn (j, (_, t2)) => | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 95 | if i = j orelse null (inter (op =) c1 t2) | 
| 35403 | 96 | then I else Int_Graph.add_edge_acyclic (i,j)) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 97 | num_branches num_branches | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 98 | end | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 99 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 100 | val default_congs = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 101 |   map (fn c => c RS eq_reflection) [@{thm "cong"}, @{thm "ext"}]
 | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 102 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 103 | (* Called on the INSTANTIATED branches of the congruence rule *) | 
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 104 | fun mk_branch ctxt t = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 105 | let | 
| 60695 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 wenzelm parents: 
59621diff
changeset | 106 | val ((params, impl), ctxt') = Variable.focus NONE t ctxt | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 107 | val (assms, concl) = Logic.strip_horn impl | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 108 | in | 
| 42495 
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
 wenzelm parents: 
42362diff
changeset | 109 | (ctxt', map #2 params, assms, rhs_of concl) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 110 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 111 | |
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 112 | fun find_cong_rule ctxt fvar h ((r,dep)::rs) t = | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 113 | (let | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 114 | val thy = Proof_Context.theory_of ctxt | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 115 | |
| 65418 
c821f1f3d92d
more general signature; works for all terms, not just frees
 Lars Hupel <lars.hupel@mytum.de> parents: 
61112diff
changeset | 116 | val tt' = Logic.mk_equals (Pattern.rewrite_term thy [(fvar, h)] [] t, t) | 
| 59582 | 117 | val (c, subs) = (Thm.concl_of r, Thm.prems_of r) | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 118 | |
| 59618 | 119 | val subst = Pattern.match thy (c, tt') (Vartab.empty, Vartab.empty) | 
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 120 | val branches = map (mk_branch ctxt o Envir.beta_norm o Envir.subst_term subst) subs | 
| 59618 | 121 | val inst = | 
| 60781 | 122 | map (fn v => (#1 v, Thm.cterm_of ctxt (Envir.subst_term subst (Var v)))) | 
| 59618 | 123 | (Term.add_vars c []) | 
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 124 | in | 
| 60781 | 125 | (infer_instantiate ctxt inst r, dep, branches) | 
| 59618 | 126 | end handle Pattern.MATCH => find_cong_rule ctxt fvar h rs t) | 
| 40317 
1eac228c52b3
replaced ancient sys_error by raise Fail, assuming that the latter is not handled specifically by the environment;
 wenzelm parents: 
36945diff
changeset | 127 | | find_cong_rule _ _ _ [] _ = raise General.Fail "No cong rule found!" | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 128 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 129 | |
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 130 | fun mk_tree fvar h ctxt t = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 131 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 132 | val congs = get_function_congs ctxt | 
| 26115 | 133 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 134 | (* FIXME: Save in theory: *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 135 | val congs_deps = map (fn c => (c, cong_deps c)) (congs @ default_congs) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 136 | |
| 65418 
c821f1f3d92d
more general signature; works for all terms, not just frees
 Lars Hupel <lars.hupel@mytum.de> parents: 
61112diff
changeset | 137 | fun matchcall (a $ b) = if a = fvar then SOME b else NONE | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 138 | | matchcall _ = NONE | 
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 139 | |
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 140 | fun mk_tree' ctxt t = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 141 | case matchcall t of | 
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 142 | SOME arg => RCall (t, mk_tree' ctxt arg) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 143 | | NONE => | 
| 65418 
c821f1f3d92d
more general signature; works for all terms, not just frees
 Lars Hupel <lars.hupel@mytum.de> parents: 
61112diff
changeset | 144 | if not (exists_subterm (fn v => v = fvar) t) then Leaf t | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 145 | else | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 146 | let | 
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 147 | val (r, dep, branches) = find_cong_rule ctxt fvar h congs_deps t | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 148 | fun subtree (ctxt', fixes, assumes, st) = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 149 | ((fixes, | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59618diff
changeset | 150 | map (Thm.assume o Thm.cterm_of ctxt) assumes), | 
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 151 | mk_tree' ctxt' st) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 152 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 153 | Cong (r, dep, map subtree branches) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 154 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 155 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 156 | mk_tree' ctxt t | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 157 | end | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 158 | |
| 59618 | 159 | fun inst_tree ctxt fvar f tr = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 160 | let | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59618diff
changeset | 161 | val cfvar = Thm.cterm_of ctxt fvar | 
| 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59618diff
changeset | 162 | val cf = Thm.cterm_of ctxt f | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 163 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 164 | fun inst_term t = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 165 | subst_bound(f, abstract_over (fvar, t)) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 166 | |
| 36945 | 167 | val inst_thm = Thm.forall_elim cf o Thm.forall_intr cfvar | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 168 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 169 | fun inst_tree_aux (Leaf t) = Leaf t | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 170 | | inst_tree_aux (Cong (crule, deps, branches)) = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 171 | Cong (inst_thm crule, deps, map inst_branch branches) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 172 | | inst_tree_aux (RCall (t, str)) = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 173 | RCall (inst_term t, inst_tree_aux str) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 174 | and inst_branch ((fxs, assms), str) = | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59618diff
changeset | 175 | ((fxs, map (Thm.assume o Thm.cterm_of ctxt o inst_term o Thm.prop_of) assms), | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 176 | inst_tree_aux str) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 177 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 178 | inst_tree_aux tr | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 179 | end | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 180 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 181 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 182 | (* Poor man's contexts: Only fixes and assumes *) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 183 | fun compose (fs1, as1) (fs2, as2) = (fs1 @ fs2, as1 @ as2) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 184 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 185 | fun export_term (fixes, assumes) = | 
| 59582 | 186 | fold_rev (curry Logic.mk_implies o Thm.prop_of) assumes | 
| 27330 | 187 | #> fold_rev (Logic.all o Free) fixes | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 188 | |
| 59618 | 189 | fun export_thm ctxt (fixes, assumes) = | 
| 59582 | 190 | fold_rev (Thm.implies_intr o Thm.cprop_of) assumes | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59618diff
changeset | 191 | #> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt o Free) fixes | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 192 | |
| 59618 | 193 | fun import_thm ctxt (fixes, athms) = | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59618diff
changeset | 194 | fold (Thm.forall_elim o Thm.cterm_of ctxt o Free) fixes | 
| 24977 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24168diff
changeset | 195 | #> fold Thm.elim_implies athms | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 196 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 197 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 198 | (* folds in the order of the dependencies of a graph. *) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 199 | fun fold_deps G f x = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 200 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 201 | fun fill_table i (T, x) = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 202 | case Inttab.lookup T i of | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 203 | SOME _ => (T, x) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 204 | | NONE => | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 205 | let | 
| 44338 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 wenzelm parents: 
42495diff
changeset | 206 | val (T', x') = Int_Graph.Keys.fold fill_table (Int_Graph.imm_succs G i) (T, x) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 207 | val (v, x'') = f (the o Inttab.lookup T') i x' | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 208 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 209 | (Inttab.update (i, v) T', x'') | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 210 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 211 | |
| 35403 | 212 | val (T, x) = fold fill_table (Int_Graph.keys G) (Inttab.empty, x) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 213 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 214 | (Inttab.fold (cons o snd) T [], x) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 215 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 216 | |
| 26115 | 217 | fun traverse_tree rcOp tr = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 218 | let | 
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 219 | fun traverse_help ctxt (Leaf _) _ x = ([], x) | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 220 | | traverse_help ctxt (RCall (t, st)) u x = | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 221 | rcOp ctxt t u (traverse_help ctxt st u x) | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 222 | | traverse_help ctxt (Cong (_, deps, branches)) u x = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 223 | let | 
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 224 | fun sub_step lu i x = | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 225 | let | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 226 | val (ctxt', subtree) = nth branches i | 
| 44338 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 wenzelm parents: 
42495diff
changeset | 227 | val used = Int_Graph.Keys.fold_rev (append o lu) (Int_Graph.imm_succs deps i) u | 
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 228 | val (subs, x') = traverse_help (compose ctxt ctxt') subtree used x | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 229 | val exported_subs = map (apfst (compose ctxt')) subs (* FIXME: Right order of composition? *) | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 230 | in | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 231 | (exported_subs, x') | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 232 | end | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 233 | in | 
| 42362 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 234 | fold_deps deps sub_step x | 
| 
5528970ac6f7
observe firm naming convention ctxt: Proof.context;
 wenzelm parents: 
42361diff
changeset | 235 | |> apfst flat | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 236 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 237 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 238 | snd o traverse_help ([], []) tr [] | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 239 | end | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 240 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
44338diff
changeset | 241 | fun rewrite_by_tree ctxt h ih x tr = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 242 | let | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59618diff
changeset | 243 | fun rewrite_help _ _ x (Leaf t) = (Thm.reflexive (Thm.cterm_of ctxt t), x) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 244 | | rewrite_help fix h_as x (RCall (_ $ arg, st)) = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 245 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 246 | val (inner, (lRi,ha)::x') = rewrite_help fix h_as x st (* "a' = a" *) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 247 | |
| 59618 | 248 | val iha = import_thm ctxt (fix, h_as) ha (* (a', h a') : G *) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 249 | |> Conv.fconv_rule (Conv.arg_conv (Conv.comb_conv (Conv.arg_conv (K inner)))) | 
| 26196 | 250 | (* (a, h a) : G *) | 
| 60801 | 251 | val inst_ih = Thm.instantiate' [] [SOME (Thm.cterm_of ctxt arg)] ih | 
| 36945 | 252 | val eq = Thm.implies_elim (Thm.implies_elim inst_ih lRi) iha (* h a = f a *) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 253 | |
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59618diff
changeset | 254 | val h_a'_eq_h_a = Thm.combination (Thm.reflexive (Thm.cterm_of ctxt h)) inner | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 255 | val h_a_eq_f_a = eq RS eq_reflection | 
| 36945 | 256 | val result = Thm.transitive h_a'_eq_h_a h_a_eq_f_a | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 257 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 258 | (result, x') | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 259 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 260 | | rewrite_help fix h_as x (Cong (crule, deps, branches)) = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 261 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 262 | fun sub_step lu i x = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 263 | let | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 264 | val ((fixes, assumes), st) = nth branches i | 
| 44338 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 wenzelm parents: 
42495diff
changeset | 265 | val used = map lu (Int_Graph.immediate_succs deps i) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 266 | |> map (fn u_eq => (u_eq RS sym) RS eq_reflection) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 267 | |> filter_out Thm.is_reflexive | 
| 26196 | 268 | |
| 59618 | 269 | val assumes' = map (simplify (put_simpset HOL_basic_ss ctxt addsimps used)) assumes | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 270 | |
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 271 | val (subeq, x') = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 272 | rewrite_help (fix @ fixes) (h_as @ assumes') x st | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 273 | val subeq_exp = | 
| 67710 
cc2db3239932
added HOLogic.mk_obj_eq convenience and eliminated some clones;
 wenzelm parents: 
67649diff
changeset | 274 | export_thm ctxt (fixes, assumes) (HOLogic.mk_obj_eq subeq) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 275 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 276 | (subeq_exp, x') | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 277 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 278 | val (subthms, x') = fold_deps deps sub_step x | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 279 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 280 | (fold_rev (curry op COMP) subthms crule, x') | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 281 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 282 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 283 | rewrite_help [] [] x tr | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 284 | end | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33519diff
changeset | 285 | |
| 19612 | 286 | end |