| author | wenzelm | 
| Tue, 31 Dec 2013 14:29:16 +0100 | |
| changeset 54883 | dd04a8b654fc | 
| parent 54742 | 7a86358a3c0b | 
| child 56245 | 84fc7dfa3cd4 | 
| permissions | -rw-r--r-- | 
| 37744 | 1 | (* Title: HOL/Tools/Function/function_lib.ML | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 2 | Author: Alexander Krauss, TU Muenchen | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 3 | |
| 41113 | 4 | Ad-hoc collection of function waiting to be eliminated, generalized, | 
| 5 | moved elsewhere or otherwise cleaned up. | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 6 | *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 7 | |
| 41113 | 8 | signature FUNCTION_LIB = | 
| 9 | sig | |
| 10 | val plural: string -> string -> 'a list -> string | |
| 11 | ||
| 12 | val dest_all_all: term -> term list * term | |
| 13 | ||
| 14 |   val map4: ('a -> 'b -> 'c -> 'd -> 'e) -> 'a list -> 'b list -> 'c list -> 'd list -> 'e list
 | |
| 15 |   val map7: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h) -> 'a list -> 'b list -> 'c list ->
 | |
| 16 | 'd list -> 'e list -> 'f list -> 'g list -> 'h list | |
| 17 | ||
| 18 |   val unordered_pairs: 'a list -> ('a * 'a) list
 | |
| 19 | ||
| 20 | val rename_bound: string -> term -> term | |
| 21 | val mk_forall_rename: (string * term) -> term -> term | |
| 22 | val forall_intr_rename: (string * cterm) -> thm -> thm | |
| 23 | ||
| 24 | datatype proof_attempt = Solved of thm | Stuck of thm | Fail | |
| 25 | val try_proof: cterm -> tactic -> proof_attempt | |
| 26 | ||
| 27 | val dest_binop_list: string -> term -> term list | |
| 28 | val regroup_conv: string -> string -> thm list -> int list -> conv | |
| 29 | val regroup_union_conv: int list -> conv | |
| 54406 | 30 | |
| 54407 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
54406diff
changeset | 31 | val inst_constrs_of: Proof.context -> typ -> term list | 
| 41113 | 32 | end | 
| 33 | ||
| 34 | structure Function_Lib: FUNCTION_LIB = | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 35 | struct | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 36 | |
| 40076 | 37 | (* "The variable" ^ plural " is" "s are" vs *) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 38 | fun plural sg pl [x] = sg | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 39 | | plural sg pl _ = pl | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 40 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 41 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 42 | (* Removes all quantifiers from a term, replacing bound variables by frees. *) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 43 | fun dest_all_all (t as (Const ("all",_) $ _)) =
 | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 44 | let | 
| 40076 | 45 | val (v,b) = Logic.dest_all t |> apfst Free | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 46 | val (vs, b') = dest_all_all b | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 47 | in | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 48 | (v :: vs, b') | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 49 | end | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 50 | | dest_all_all t = ([],t) | 
| 33611 | 51 | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 52 | |
| 54406 | 53 | fun map4 f = Ctr_Sugar_Util.map4 f | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 54 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 55 | fun map7 _ [] [] [] [] [] [] [] = [] | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 56 | | map7 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (w :: ws) (b :: bs) = f x y z u v w b :: map7 f xs ys zs us vs ws bs | 
| 40722 
441260986b63
make two copies (!) of Library.UnequalLengths coincide with ListPair.UnequalLengths;
 wenzelm parents: 
40076diff
changeset | 57 | | map7 _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths; | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 58 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 59 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 60 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 61 | (* forms all "unordered pairs": [1, 2, 3] ==> [(1, 1), (1, 2), (1, 3), (2, 2), (2, 3), (3, 3)] *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 62 | fun unordered_pairs [] = [] | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 63 | | unordered_pairs (x::xs) = map (pair x) (x::xs) @ unordered_pairs xs | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 64 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 65 | |
| 42497 
89acb654d8eb
inlined Function_Lib.replace_frees, which is used only once
 krauss parents: 
42483diff
changeset | 66 | (* renaming bound variables *) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 67 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 68 | fun rename_bound n (Q $ Abs (_, T, b)) = (Q $ Abs (n, T, b)) | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 69 | | rename_bound n _ = raise Match | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 70 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 71 | fun mk_forall_rename (n, v) = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 72 | rename_bound n o Logic.all v | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 73 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 74 | fun forall_intr_rename (n, cv) thm = | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 75 | let | 
| 36945 | 76 | val allthm = Thm.forall_intr cv thm | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 77 | val (_ $ abs) = prop_of allthm | 
| 45156 | 78 | in Thm.rename_boundvars abs (Abs (n, dummyT, Term.dummy)) allthm end | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 79 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 80 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 81 | datatype proof_attempt = Solved of thm | Stuck of thm | Fail | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 82 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 83 | fun try_proof cgoal tac = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 84 | case SINGLE tac (Goal.init cgoal) of | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 85 | NONE => Fail | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 86 | | SOME st => | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 87 | if Thm.no_prems st | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 88 | then Solved (Goal.finish (Syntax.init_pretty_global (Thm.theory_of_cterm cgoal)) st) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 89 | else Stuck st | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 90 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 91 | |
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 92 | fun dest_binop_list cn (t as (Const (n, _) $ a $ b)) = | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 93 | if cn = n then dest_binop_list cn a @ dest_binop_list cn b else [ t ] | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 94 | | dest_binop_list _ t = [ t ] | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 95 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 96 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 97 | (* separate two parts in a +-expression: | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 98 | "a + b + c + d + e" --> "(a + b + d) + (c + e)" | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 99 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 100 | Here, + can be any binary operation that is AC. | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 101 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 102 |    cn - The name of the binop-constructor (e.g. @{const_name Un})
 | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 103 | ac - the AC rewrite rules for cn | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 104 | is - the list of indices of the expressions that should become the first part | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 105 | (e.g. [0,1,3] in the above example) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 106 | *) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 107 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 108 | fun regroup_conv neu cn ac is ct = | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 109 | let | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54407diff
changeset | 110 | val thy = theory_of_cterm ct | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54407diff
changeset | 111 | val ctxt = Proof_Context.init_global thy (* FIXME proper context!? *) | 
| 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54407diff
changeset | 112 | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 113 | val mk = HOLogic.mk_binop cn | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 114 | val t = term_of ct | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 115 | val xs = dest_binop_list cn t | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 116 | val js = subtract (op =) is (0 upto (length xs) - 1) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 117 | val ty = fastype_of t | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 118 | in | 
| 54883 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 wenzelm parents: 
54742diff
changeset | 119 | Goal.prove_internal ctxt [] | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 120 | (cterm_of thy | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 121 | (Logic.mk_equals (t, | 
| 33611 | 122 | if null is | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 123 | then mk (Const (neu, ty), foldr1 mk (map (nth xs) js)) | 
| 33611 | 124 | else if null js | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 125 | then mk (foldr1 mk (map (nth xs) is), Const (neu, ty)) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 126 | else mk (foldr1 mk (map (nth xs) is), foldr1 mk (map (nth xs) js))))) | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54407diff
changeset | 127 | (K (rewrite_goals_tac ctxt ac | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 128 | THEN rtac Drule.reflexive_thm 1)) | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 129 | end | 
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 130 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 131 | (* instance for unions *) | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 132 | val regroup_union_conv = | 
| 35402 | 133 |   regroup_conv @{const_abbrev Set.empty} @{const_name Lattices.sup}
 | 
| 34232 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 134 | (map (fn t => t RS eq_reflection) | 
| 
36a2a3029fd3
new year's resolution: reindented code in function package
 krauss parents: 
33855diff
changeset | 135 |       (@{thms Un_ac} @ @{thms Un_empty_right} @ @{thms Un_empty_left}))
 | 
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 136 | |
| 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 137 | |
| 54407 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
54406diff
changeset | 138 | fun inst_constrs_of ctxt (Type (name, Ts)) = | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
54406diff
changeset | 139 | map (Ctr_Sugar.mk_ctr Ts) (#ctrs (the (Ctr_Sugar.ctr_sugar_of ctxt name))) | 
| 
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
 blanchet parents: 
54406diff
changeset | 140 | | inst_constrs_of _ _ = raise Match | 
| 54406 | 141 | |
| 33099 
b8cdd3d73022
function package: more standard names for structures and files
 krauss parents: diff
changeset | 142 | end |