author | wenzelm |
Fri, 26 Jul 2019 09:59:11 +0200 | |
changeset 70415 | 3c20a86f14f1 |
parent 69593 | 3dda49e08b9d |
child 74525 | c960bfcb91db |
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 |
|
62093 | 10 |
val function_internals: bool Config.T |
61121
efe8b18306b7
do not expose low-level "_def" facts of 'function' definitions, to avoid potential confusion with the situation of plain 'definition';
wenzelm
parents:
60328
diff
changeset
|
11 |
|
63005 | 12 |
val derived_name: binding -> string -> binding |
13 |
val derived_name_suffix: binding -> string -> binding |
|
14 |
||
41113 | 15 |
val plural: string -> string -> 'a list -> string |
16 |
||
17 |
val dest_all_all: term -> term list * term |
|
18 |
||
19 |
val unordered_pairs: 'a list -> ('a * 'a) list |
|
20 |
||
21 |
val rename_bound: string -> term -> term |
|
22 |
val mk_forall_rename: (string * term) -> term -> term |
|
23 |
val forall_intr_rename: (string * cterm) -> thm -> thm |
|
24 |
||
25 |
datatype proof_attempt = Solved of thm | Stuck of thm | Fail |
|
60328 | 26 |
val try_proof: Proof.context -> cterm -> tactic -> proof_attempt |
41113 | 27 |
|
28 |
val dest_binop_list: string -> term -> term list |
|
59625 | 29 |
val regroup_conv: Proof.context -> string -> string -> thm list -> int list -> conv |
30 |
val regroup_union_conv: Proof.context -> int list -> conv |
|
54406 | 31 |
|
54407
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
blanchet
parents:
54406
diff
changeset
|
32 |
val inst_constrs_of: Proof.context -> typ -> term list |
41113 | 33 |
end |
34 |
||
35 |
structure Function_Lib: FUNCTION_LIB = |
|
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
36 |
struct |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
37 |
|
69593 | 38 |
val function_internals = Attrib.setup_config_bool \<^binding>\<open>function_internals\<close> (K false) |
61121
efe8b18306b7
do not expose low-level "_def" facts of 'function' definitions, to avoid potential confusion with the situation of plain 'definition';
wenzelm
parents:
60328
diff
changeset
|
39 |
|
63005 | 40 |
fun derived_name binding name = |
41 |
Binding.reset_pos (Binding.qualify_name true binding name) |
|
42 |
||
63011 | 43 |
fun derived_name_suffix binding sffx = |
44 |
Binding.reset_pos (Binding.map_name (suffix sffx) binding) |
|
63005 | 45 |
|
61121
efe8b18306b7
do not expose low-level "_def" facts of 'function' definitions, to avoid potential confusion with the situation of plain 'definition';
wenzelm
parents:
60328
diff
changeset
|
46 |
|
40076 | 47 |
(* "The variable" ^ plural " is" "s are" vs *) |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
48 |
fun plural sg pl [x] = sg |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
49 |
| plural sg pl _ = pl |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
50 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
51 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
52 |
(* Removes all quantifiers from a term, replacing bound variables by frees. *) |
67149 | 53 |
fun dest_all_all (t as (Const (\<^const_name>\<open>Pure.all\<close>,_) $ _)) = |
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33855
diff
changeset
|
54 |
let |
40076 | 55 |
val (v,b) = Logic.dest_all t |> apfst Free |
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33855
diff
changeset
|
56 |
val (vs, b') = dest_all_all b |
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33855
diff
changeset
|
57 |
in |
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33855
diff
changeset
|
58 |
(v :: vs, b') |
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33855
diff
changeset
|
59 |
end |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
60 |
| dest_all_all t = ([],t) |
33611 | 61 |
|
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
62 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
63 |
(* 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
|
64 |
fun unordered_pairs [] = [] |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
65 |
| 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
|
66 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
67 |
|
42497
89acb654d8eb
inlined Function_Lib.replace_frees, which is used only once
krauss
parents:
42483
diff
changeset
|
68 |
(* renaming bound variables *) |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
69 |
|
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33855
diff
changeset
|
70 |
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
|
71 |
| rename_bound n _ = raise Match |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
72 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
73 |
fun mk_forall_rename (n, v) = |
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33855
diff
changeset
|
74 |
rename_bound n o Logic.all v |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
75 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
76 |
fun forall_intr_rename (n, cv) thm = |
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33855
diff
changeset
|
77 |
let |
36945 | 78 |
val allthm = Thm.forall_intr cv thm |
59582 | 79 |
val (_ $ abs) = Thm.prop_of allthm |
45156 | 80 |
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
|
81 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
82 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
83 |
datatype proof_attempt = Solved of thm | Stuck of thm | Fail |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
84 |
|
60328 | 85 |
fun try_proof ctxt cgoal tac = |
86 |
(case SINGLE tac (Goal.init cgoal) of |
|
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33855
diff
changeset
|
87 |
NONE => Fail |
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33855
diff
changeset
|
88 |
| SOME st => |
60328 | 89 |
if Thm.no_prems st |
90 |
then Solved (Goal.finish ctxt st) |
|
91 |
else Stuck st) |
|
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
92 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
93 |
|
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33855
diff
changeset
|
94 |
fun dest_binop_list cn (t as (Const (n, _) $ a $ b)) = |
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33855
diff
changeset
|
95 |
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
|
96 |
| dest_binop_list _ t = [ t ] |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
97 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
98 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
99 |
(* separate two parts in a +-expression: |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
100 |
"a + b + c + d + e" --> "(a + b + d) + (c + e)" |
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 |
Here, + can be any binary operation that is AC. |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
103 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
104 |
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
|
105 |
ac - the AC rewrite rules for cn |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
106 |
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
|
107 |
(e.g. [0,1,3] in the above example) |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
108 |
*) |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
109 |
|
59625 | 110 |
fun regroup_conv ctxt neu cn ac is ct = |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
111 |
let |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
112 |
val mk = HOLogic.mk_binop cn |
59582 | 113 |
val t = Thm.term_of ct |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
114 |
val xs = dest_binop_list cn t |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
115 |
val js = subtract (op =) is (0 upto (length xs) - 1) |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
116 |
val ty = fastype_of t |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
117 |
in |
54883
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
54742
diff
changeset
|
118 |
Goal.prove_internal ctxt [] |
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59618
diff
changeset
|
119 |
(Thm.cterm_of ctxt |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
120 |
(Logic.mk_equals (t, |
33611 | 121 |
if null is |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
122 |
then mk (Const (neu, ty), foldr1 mk (map (nth xs) js)) |
33611 | 123 |
else if null js |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
124 |
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
|
125 |
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:
54407
diff
changeset
|
126 |
(K (rewrite_goals_tac ctxt ac |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58839
diff
changeset
|
127 |
THEN resolve_tac ctxt [Drule.reflexive_thm] 1)) |
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
128 |
end |
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
129 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
130 |
(* instance for unions *) |
59625 | 131 |
fun regroup_union_conv ctxt = |
67149 | 132 |
regroup_conv ctxt \<^const_abbrev>\<open>Set.empty\<close> \<^const_name>\<open>Lattices.sup\<close> |
34232
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33855
diff
changeset
|
133 |
(map (fn t => t RS eq_reflection) |
36a2a3029fd3
new year's resolution: reindented code in function package
krauss
parents:
33855
diff
changeset
|
134 |
(@{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
|
135 |
|
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
136 |
|
54407
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
blanchet
parents:
54406
diff
changeset
|
137 |
fun inst_constrs_of ctxt (Type (name, Ts)) = |
e95831757903
ported part of function package to new 'Ctr_Sugar' abstraction
blanchet
parents:
54406
diff
changeset
|
138 |
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:
54406
diff
changeset
|
139 |
| inst_constrs_of _ _ = raise Match |
54406 | 140 |
|
33099
b8cdd3d73022
function package: more standard names for structures and files
krauss
parents:
diff
changeset
|
141 |
end |