| author | paulson | 
| Mon, 14 Dec 2009 11:30:13 +0000 | |
| changeset 34087 | c907edcaab36 | 
| parent 33519 | e31a85f92ce9 | 
| child 34232 | 36a2a3029fd3 | 
| permissions | -rw-r--r-- | 
| 31775 | 1 | (* Title: HOL/Tools/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 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 4 | A package for general recursive function definitions. | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 5 | Builds and traverses trees of nested contexts along a term. | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 6 | *) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 7 | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33049diff
changeset | 8 | signature FUNCTION_CTXTREE = | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 9 | sig | 
| 26115 | 10 | type ctxt = (string * typ) list * thm list (* poor man's contexts: fixes + assumes *) | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 11 | type ctx_tree | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 12 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 13 | (* FIXME: This interface is a mess and needs to be cleaned up! *) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33049diff
changeset | 14 | val get_function_congs : Proof.context -> thm list | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33049diff
changeset | 15 | val add_function_cong : thm -> Context.generic -> Context.generic | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33049diff
changeset | 16 | val map_function_congs : (thm list -> thm list) -> Context.generic -> Context.generic | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 17 | |
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 18 | val cong_add: attribute | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 19 | val cong_del: attribute | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 20 | |
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 21 | val mk_tree: (string * typ) -> term -> Proof.context -> term -> ctx_tree | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 22 | |
| 23819 | 23 | val inst_tree: theory -> term -> term -> ctx_tree -> ctx_tree | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 24 | |
| 26115 | 25 | val export_term : ctxt -> term -> term | 
| 26 | val export_thm : theory -> ctxt -> thm -> thm | |
| 27 | val import_thm : theory -> ctxt -> thm -> thm | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 28 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 29 | val traverse_tree : | 
| 26115 | 30 | (ctxt -> term -> | 
| 31 | (ctxt * thm) list -> | |
| 32 | (ctxt * thm) list * 'b -> | |
| 33 | (ctxt * thm) list * 'b) | |
| 23819 | 34 | -> ctx_tree -> 'b -> 'b | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 35 | |
| 23819 | 36 | val rewrite_by_tree : theory -> term -> thm -> (thm * thm) list -> ctx_tree -> thm * (thm * thm) list | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 37 | end | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 38 | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33049diff
changeset | 39 | structure Function_Ctx_Tree : FUNCTION_CTXTREE = | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 40 | struct | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 41 | |
| 26115 | 42 | type ctxt = (string * typ) list * thm list | 
| 43 | ||
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33049diff
changeset | 44 | open Function_Common | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33049diff
changeset | 45 | open Function_Lib | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 46 | |
| 33519 | 47 | structure FunctionCongs = Generic_Data | 
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 48 | ( | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 49 | type T = thm list | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 50 | val empty = [] | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 51 | val extend = I | 
| 33519 | 52 | val merge = Thm.merge_thms | 
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 53 | ); | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 54 | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33049diff
changeset | 55 | val get_function_congs = FunctionCongs.get o Context.Proof | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33049diff
changeset | 56 | val map_function_congs = FunctionCongs.map | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33049diff
changeset | 57 | val add_function_cong = FunctionCongs.map o Thm.add_thm | 
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 58 | |
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 59 | (* congruence rules *) | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 60 | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33049diff
changeset | 61 | val cong_add = Thm.declaration_attribute (map_function_congs o Thm.add_thm o safe_mk_meta_eq); | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33049diff
changeset | 62 | val cong_del = Thm.declaration_attribute (map_function_congs 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 | 63 | |
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 64 | |
| 23819 | 65 | type depgraph = int IntGraph.T | 
| 66 | ||
| 67 | datatype ctx_tree | |
| 68 | = Leaf of term | |
| 26115 | 69 | | Cong of (thm * depgraph * (ctxt * ctx_tree) list) | 
| 23819 | 70 | | RCall of (term * ctx_tree) | 
| 71 | ||
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 72 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 73 | (* Maps "Trueprop A = B" to "A" *) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 74 | 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 | 75 | |
| 
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 | (*** Dependency analysis for congruence rules ***) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 78 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 79 | fun branch_vars t = | 
| 21237 | 80 | let | 
| 21188 | 81 | val t' = snd (dest_all_all t) | 
| 26529 | 82 | val (assumes, concl) = Logic.strip_horn t' | 
| 29329 | 83 | in (fold Term.add_vars assumes [], Term.add_vars concl []) | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 84 | end | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 85 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 86 | fun cong_deps crule = | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 87 | let | 
| 26115 | 88 | val num_branches = map_index (apsnd branch_vars) (prems_of crule) | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 89 | in | 
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 90 | IntGraph.empty | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 91 | |> fold (fn (i,_)=> IntGraph.new_node (i,i)) num_branches | 
| 29329 | 92 | |> fold_product (fn (i, (c1, _)) => fn (j, (_, t2)) => | 
| 33049 
c38f02fdf35d
curried inter as canonical list operation (beware of argument order)
 haftmann parents: 
33038diff
changeset | 93 | if i = j orelse null (inter (op =) c1 t2) | 
| 29329 | 94 | then I else IntGraph.add_edge_acyclic (i,j)) | 
| 25538 | 95 | num_branches num_branches | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 96 | end | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 97 | |
| 26115 | 98 | val default_congs = 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 | 99 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 100 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 101 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 102 | (* Called on the INSTANTIATED branches of the congruence rule *) | 
| 19922 | 103 | fun mk_branch ctx t = | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 104 | let | 
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 105 | val (ctx', fixes, impl) = dest_all_all_ctx ctx t | 
| 26529 | 106 | val (assms, concl) = Logic.strip_horn impl | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 107 | in | 
| 26529 | 108 | (ctx', fixes, assms, rhs_of concl) | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 109 | end | 
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 110 | |
| 21100 | 111 | fun find_cong_rule ctx fvar h ((r,dep)::rs) t = | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 112 | (let | 
| 21100 | 113 | val thy = ProofContext.theory_of ctx | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 114 | |
| 21100 | 115 | val tt' = Logic.mk_equals (Pattern.rewrite_term thy [(Free fvar, h)] [] t, t) | 
| 116 | val (c, subs) = (concl_of r, prems_of r) | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 117 | |
| 21100 | 118 | val subst = Pattern.match (ProofContext.theory_of ctx) (c, tt') (Vartab.empty, Vartab.empty) | 
| 32035 | 119 | val branches = map (mk_branch ctx o Envir.beta_norm o Envir.subst_term subst) subs | 
| 120 | val inst = map (fn v => | |
| 121 | (cterm_of thy (Var v), cterm_of thy (Envir.subst_term subst (Var v)))) (Term.add_vars c []) | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 122 | in | 
| 21237 | 123 | (cterm_instantiate inst r, dep, branches) | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 124 | end | 
| 21100 | 125 | handle Pattern.MATCH => find_cong_rule ctx fvar h rs t) | 
| 31775 | 126 | | find_cong_rule _ _ _ [] _ = sys_error "Function/context_tree.ML: No cong rule found!" | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 127 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 128 | |
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 129 | fun mk_tree fvar h ctxt t = | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 130 | let | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: 
33049diff
changeset | 131 | val congs = get_function_congs ctxt | 
| 26115 | 132 | val congs_deps = map (fn c => (c, cong_deps c)) (congs @ default_congs) (* FIXME: Save in theory *) | 
| 133 | ||
| 134 | fun matchcall (a $ b) = if a = Free fvar then SOME b else NONE | |
| 135 | | matchcall _ = NONE | |
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 136 | |
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 137 | fun mk_tree' ctx t = | 
| 26115 | 138 | case matchcall t of | 
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 139 | SOME arg => RCall (t, mk_tree' ctx arg) | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 140 | | NONE => | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 141 | if not (exists_subterm (fn Free v => v = fvar | _ => false) t) then Leaf t | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 142 | else | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 143 | let val (r, dep, branches) = find_cong_rule ctx fvar h congs_deps t in | 
| 26115 | 144 | Cong (r, dep, | 
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 145 | map (fn (ctx', fixes, assumes, st) => | 
| 26115 | 146 | ((fixes, map (assume o cterm_of (ProofContext.theory_of ctx)) assumes), | 
| 24168 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 147 | mk_tree' ctx' st)) branches) | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 148 | end | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 149 | in | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 150 | mk_tree' ctxt t | 
| 
86a03a092062
simplified internal interfaces; cong rules are now handled directly by "context_tree.ML"
 krauss parents: 
23819diff
changeset | 151 | end | 
| 21237 | 152 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 153 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 154 | fun inst_tree thy fvar f tr = | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 155 | let | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 156 | val cfvar = cterm_of thy fvar | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 157 | val cf = cterm_of thy f | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 158 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 159 | fun inst_term t = | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 160 | subst_bound(f, abstract_over (fvar, t)) | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 161 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 162 | val inst_thm = forall_elim cf o forall_intr cfvar | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 163 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 164 | fun inst_tree_aux (Leaf t) = Leaf t | 
| 26115 | 165 | | inst_tree_aux (Cong (crule, deps, branches)) = | 
| 166 | Cong (inst_thm crule, deps, map inst_branch branches) | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 167 | | inst_tree_aux (RCall (t, str)) = | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 168 | RCall (inst_term t, inst_tree_aux str) | 
| 26115 | 169 | and inst_branch ((fxs, assms), str) = | 
| 170 | ((fxs, map (assume o cterm_of thy o inst_term o prop_of) assms), inst_tree_aux str) | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 171 | in | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 172 | inst_tree_aux tr | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 173 | end | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 174 | |
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 175 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 176 | (* 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 | 177 | 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 | 178 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 179 | fun export_term (fixes, assumes) = | 
| 26115 | 180 | fold_rev (curry Logic.mk_implies o prop_of) assumes | 
| 27330 | 181 | #> 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 | 182 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 183 | fun export_thm thy (fixes, assumes) = | 
| 26115 | 184 | fold_rev (implies_intr o cprop_of) assumes | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 185 | #> fold_rev (forall_intr o cterm_of thy o Free) fixes | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 186 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 187 | fun import_thm thy (fixes, athms) = | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 188 | fold (forall_elim o cterm_of thy o Free) fixes | 
| 24977 
9f98751c9628
replaced (flip Thm.implies_elim) by Thm.elim_implies;
 wenzelm parents: 
24168diff
changeset | 189 | #> fold Thm.elim_implies athms | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 190 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 191 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 192 | (* 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 | 193 | fun fold_deps G f x = | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 194 | let | 
| 26115 | 195 | fun fill_table i (T, x) = | 
| 196 | case Inttab.lookup T i of | |
| 197 | SOME _ => (T, x) | |
| 198 | | NONE => | |
| 199 | let | |
| 200 | val (T', x') = fold fill_table (IntGraph.imm_succs G i) (T, x) | |
| 201 | val (v, x'') = f (the o Inttab.lookup T') i x' | |
| 202 | in | |
| 203 | (Inttab.update (i, v) T', x'') | |
| 204 | end | |
| 205 | ||
| 206 | val (T, x) = fold fill_table (IntGraph.keys G) (Inttab.empty, x) | |
| 21237 | 207 | in | 
| 26115 | 208 | (Inttab.fold (cons o snd) T [], x) | 
| 21237 | 209 | end | 
| 26115 | 210 | |
| 211 | fun traverse_tree rcOp tr = | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 212 | let | 
| 26115 | 213 | fun traverse_help ctx (Leaf _) _ x = ([], x) | 
| 21237 | 214 | | traverse_help ctx (RCall (t, st)) u x = | 
| 215 | rcOp ctx t u (traverse_help ctx st u x) | |
| 26115 | 216 | | traverse_help ctx (Cong (_, deps, branches)) u x = | 
| 21237 | 217 | let | 
| 218 | fun sub_step lu i x = | |
| 219 | let | |
| 26115 | 220 | val (ctx', subtree) = nth branches i | 
| 221 | val used = fold_rev (append o lu) (IntGraph.imm_succs deps i) u | |
| 222 | val (subs, x') = traverse_help (compose ctx ctx') subtree used x | |
| 223 | val exported_subs = map (apfst (compose ctx')) subs (* FIXME: Right order of composition? *) | |
| 21237 | 224 | in | 
| 26115 | 225 | (exported_subs, x') | 
| 21237 | 226 | end | 
| 227 | in | |
| 26115 | 228 | fold_deps deps sub_step x | 
| 229 | |> apfst flat | |
| 21237 | 230 | end | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 231 | in | 
| 26115 | 232 | snd o traverse_help ([], []) tr [] | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 233 | end | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 234 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20289diff
changeset | 235 | fun rewrite_by_tree thy h ih x tr = | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 236 | let | 
| 26196 | 237 | fun rewrite_help _ _ x (Leaf t) = (reflexive (cterm_of thy t), x) | 
| 238 | | rewrite_help fix h_as x (RCall (_ $ arg, st)) = | |
| 21237 | 239 | let | 
| 26196 | 240 | val (inner, (lRi,ha)::x') = rewrite_help fix h_as x st (* "a' = a" *) | 
| 21237 | 241 | |
| 242 | val iha = import_thm thy (fix, h_as) ha (* (a', h a') : G *) | |
| 26196 | 243 | |> Conv.fconv_rule (Conv.arg_conv (Conv.comb_conv (Conv.arg_conv (K inner)))) | 
| 244 | (* (a, h a) : G *) | |
| 21237 | 245 | val inst_ih = instantiate' [] [SOME (cterm_of thy arg)] ih | 
| 26196 | 246 | val eq = implies_elim (implies_elim inst_ih lRi) iha (* h a = f a *) | 
| 21237 | 247 | |
| 26196 | 248 | val h_a'_eq_h_a = combination (reflexive (cterm_of thy h)) inner | 
| 249 | val h_a_eq_f_a = eq RS eq_reflection | |
| 250 | val result = transitive h_a'_eq_h_a h_a_eq_f_a | |
| 21237 | 251 | in | 
| 252 | (result, x') | |
| 253 | end | |
| 26196 | 254 | | rewrite_help fix h_as x (Cong (crule, deps, branches)) = | 
| 21237 | 255 | let | 
| 256 | fun sub_step lu i x = | |
| 257 | let | |
| 26115 | 258 | val ((fixes, assumes), st) = nth branches i | 
| 26196 | 259 | val used = map lu (IntGraph.imm_succs deps i) | 
| 260 | |> map (fn u_eq => (u_eq RS sym) RS eq_reflection) | |
| 261 | |> filter_out Thm.is_reflexive | |
| 262 | ||
| 263 | val assumes' = map (simplify (HOL_basic_ss addsimps used)) assumes | |
| 21237 | 264 | |
| 26196 | 265 | val (subeq, x') = rewrite_help (fix @ fixes) (h_as @ assumes') x st | 
| 26115 | 266 | val subeq_exp = export_thm thy (fixes, assumes) (subeq RS meta_eq_to_obj_eq) | 
| 21237 | 267 | in | 
| 268 | (subeq_exp, x') | |
| 269 | end | |
| 270 | ||
| 271 | val (subthms, x') = fold_deps deps sub_step x | |
| 272 | in | |
| 273 | (fold_rev (curry op COMP) subthms crule, x') | |
| 274 | end | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 275 | in | 
| 26196 | 276 | rewrite_help [] [] x tr | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 277 | end | 
| 21237 | 278 | |
| 19612 | 279 | end |