author | nipkow |
Fri, 25 Oct 2024 16:57:17 +0200 | |
changeset 81255 | 47530e9a7c33 |
parent 81147 | 503e5280ba72 |
child 81289 | 1eddc169baea |
permissions | -rw-r--r-- |
79490 | 1 |
|
2 |
signature TIMING_FUNCTIONS = |
|
3 |
sig |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
4 |
type 'a wctxt = { |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
5 |
ctxt: local_theory, |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
6 |
origins: term list, |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
7 |
f: term -> 'a |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
8 |
} |
79490 | 9 |
type 'a converter = { |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
10 |
constc : 'a wctxt -> term -> 'a, |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
11 |
funcc : 'a wctxt -> term -> term list -> 'a, |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
12 |
ifc : 'a wctxt -> typ -> term -> term -> term -> 'a, |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
13 |
casec : 'a wctxt -> term -> term list -> 'a, |
81147 | 14 |
letc : 'a wctxt -> typ -> term -> (string * typ) list -> term -> 'a |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
15 |
} |
79494
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
16 |
val walk : local_theory -> term list -> 'a converter -> term -> 'a |
81147 | 17 |
val Iconst : term wctxt -> term -> term |
18 |
val Ifunc : term wctxt -> term -> term list -> term |
|
19 |
val Iif : term wctxt -> typ -> term -> term -> term -> term |
|
20 |
val Icase : term wctxt -> term -> term list -> term |
|
21 |
val Ilet : term wctxt -> typ -> term -> (string * typ) list -> term -> term |
|
79490 | 22 |
|
79494
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
23 |
type pfunc = { names : string list, terms : term list, typs : typ list } |
79490 | 24 |
val fun_pretty': Proof.context -> pfunc -> Pretty.T |
25 |
val fun_pretty: Proof.context -> Function.info -> Pretty.T |
|
26 |
val print_timing': Proof.context -> pfunc -> pfunc -> unit |
|
27 |
val print_timing: Proof.context -> Function.info -> Function.info -> unit |
|
28 |
||
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
29 |
val reg_and_proove_time_func: local_theory -> term list -> term list |
81147 | 30 |
-> bool -> bool -> Function.info * local_theory |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
31 |
val reg_time_func: local_theory -> term list -> term list |
81147 | 32 |
-> bool -> bool -> Function.info * local_theory |
79490 | 33 |
|
34 |
val time_dom_tac: Proof.context -> thm -> thm list -> int -> tactic |
|
35 |
||
36 |
end |
|
37 |
||
38 |
structure Timing_Functions : TIMING_FUNCTIONS = |
|
39 |
struct |
|
40 |
(* Configure config variable to adjust the prefix *) |
|
41 |
val bprefix = Attrib.setup_config_string @{binding "time_prefix"} (K "T_") |
|
81147 | 42 |
val bprefix_snd = Attrib.setup_config_string @{binding "time_prefix_snd"} (K "T2_") |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
43 |
(* Configure config variable to adjust the suffix *) |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
44 |
val bsuffix = Attrib.setup_config_string @{binding "time_suffix"} (K "") |
79490 | 45 |
|
46 |
(* some default values to build terms easier *) |
|
47 |
val zero = Const (@{const_name "Groups.zero"}, HOLogic.natT) |
|
48 |
val one = Const (@{const_name "Groups.one"}, HOLogic.natT) |
|
49 |
(* Extracts terms from function info *) |
|
50 |
fun terms_of_info (info: Function.info) = |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
51 |
map Thm.prop_of (case #simps info of SOME s => s |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
52 |
| NONE => error "No terms of function found in info") |
79490 | 53 |
|
54 |
type pfunc = { |
|
79494
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
55 |
names : string list, |
79490 | 56 |
terms : term list, |
79494
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
57 |
typs : typ list |
79490 | 58 |
} |
59 |
fun info_pfunc (info: Function.info): pfunc = |
|
60 |
let |
|
61 |
val {defname, fs, ...} = info; |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
62 |
val T = case hd fs of (Const (_,T)) => T |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
63 |
| (Free (_,T)) => T |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
64 |
| _ => error "Internal error: Invalid info to print" |
79490 | 65 |
in |
79494
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
66 |
{ names=[Binding.name_of defname], terms=terms_of_info info, typs=[T] } |
79490 | 67 |
end |
68 |
||
69 |
(* Auxiliary functions for printing functions *) |
|
70 |
fun fun_pretty' ctxt (pfunc: pfunc) = |
|
71 |
let |
|
79494
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
72 |
val {names, terms, typs} = pfunc; |
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
73 |
val header_beg = Pretty.str "fun "; |
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
74 |
fun prepHeadCont (nm,T) = [Pretty.str (nm ^ " :: "), (Pretty.quote (Syntax.pretty_typ ctxt T))] |
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
75 |
val header_content = |
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
76 |
List.concat (prepHeadCont (hd names,hd typs) :: map ((fn l => Pretty.str "\nand " :: l) o prepHeadCont) (ListPair.zip (tl names, tl typs))); |
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
77 |
val header_end = Pretty.str " where\n "; |
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
78 |
val header = [header_beg] @ header_content @ [header_end]; |
79490 | 79 |
fun separate sep prts = |
80 |
flat (Library.separate [Pretty.str sep] (map single prts)); |
|
81 |
val ptrms = (separate "\n| " (map (Syntax.pretty_term ctxt) terms)); |
|
82 |
in |
|
83 |
Pretty.text_fold (header @ ptrms) |
|
84 |
end |
|
85 |
fun fun_pretty ctxt = fun_pretty' ctxt o info_pfunc |
|
86 |
fun print_timing' ctxt (opfunc: pfunc) (tpfunc: pfunc) = |
|
87 |
let |
|
79494
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
88 |
val {names, ...} = opfunc; |
79490 | 89 |
val poriginal = Pretty.item [Pretty.str "Original function:\n", fun_pretty' ctxt opfunc] |
90 |
val ptiming = Pretty.item [Pretty.str ("Running time function:\n"), fun_pretty' ctxt tpfunc] |
|
91 |
in |
|
81147 | 92 |
Pretty.writeln (Pretty.text_fold [ |
93 |
Pretty.str ("Converting " ^ (hd names) ^ (String.concat (map (fn nm => ", " ^ nm) (tl names))) ^ "\n"), |
|
94 |
poriginal, Pretty.str "\n", ptiming]) |
|
79490 | 95 |
end |
96 |
fun print_timing ctxt (oinfo: Function.info) (tinfo: Function.info) = |
|
97 |
print_timing' ctxt (info_pfunc oinfo) (info_pfunc tinfo) |
|
98 |
||
81147 | 99 |
fun print_lemma ctxt defs (T_terms: term list) = |
100 |
let |
|
101 |
val names = |
|
102 |
defs |
|
103 |
|> map snd |
|
104 |
|> map (fn s => "_" ^ s) |
|
105 |
|> List.foldr (op ^) "" |
|
106 |
val begin = "lemma T" ^ names ^ "_simps [simp,code]:\n" |
|
107 |
fun convLine T_term = |
|
108 |
" \"" ^ Syntax.string_of_term ctxt T_term ^ "\"\n" |
|
109 |
val lines = map convLine T_terms |
|
110 |
fun convDefs def = " " ^ (fst def) |
|
111 |
val proof = " by (simp_all add:" :: (map convDefs defs) @ [")"] |
|
112 |
val _ = Pretty.writeln (Pretty.str "Characteristic recursion equations can be derived:") |
|
113 |
in |
|
114 |
(begin :: lines @ proof) |
|
115 |
|> String.concat |
|
116 |
(* |> Active.sendback_markup_properties [Markup.padding_fun] *) |
|
117 |
|> Pretty.str |
|
118 |
|> Pretty.writeln |
|
119 |
end |
|
120 |
||
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
121 |
fun contains l e = exists (fn e' => e' = e) l |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
122 |
fun contains' comp l e = exists (comp e) l |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
123 |
(* Split name by . *) |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
124 |
val split_name = String.fields (fn s => s = #".") |
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
125 |
|
79490 | 126 |
(* returns true if it's an if term *) |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
127 |
fun is_if (Const (@{const_name "HOL.If"},_)) = true |
79490 | 128 |
| is_if _ = false |
129 |
(* returns true if it's a case term *) |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
130 |
fun is_case (Const (n,_)) = n |> split_name |> List.last |> String.isPrefix "case_" |
79490 | 131 |
| is_case _ = false |
132 |
(* returns true if it's a let term *) |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
133 |
fun is_let (Const (@{const_name "HOL.Let"},_)) = true |
79490 | 134 |
| is_let _ = false |
135 |
(* change type of original function to new type (_ \<Rightarrow> ... \<Rightarrow> _ to _ \<Rightarrow> ... \<Rightarrow> nat) |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
136 |
and replace all function arguments f with (t*T_f) if used *) |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
137 |
fun change_typ' used (Type ("fun", [T1, T2])) = |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
138 |
Type ("fun", [check_for_fun' (used 0) T1, change_typ' (fn i => used (i+1)) T2]) |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
139 |
| change_typ' _ _ = HOLogic.natT |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
140 |
and check_for_fun' true (f as Type ("fun", [_,_])) = HOLogic.mk_prodT (f, change_typ' (K false) f) |
81147 | 141 |
| check_for_fun' false (f as Type ("fun", [_,_])) = change_typ' (K true) f |
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
142 |
| check_for_fun' _ t = t |
81147 | 143 |
val change_typ = change_typ' (K true) |
79490 | 144 |
(* Convert string name of function to its timing equivalent *) |
81147 | 145 |
fun fun_name_to_time' ctxt s second name = |
79490 | 146 |
let |
81147 | 147 |
val prefix = Config.get ctxt (if second then bprefix_snd else bprefix) |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
148 |
val suffix = (if s then Config.get ctxt bsuffix else "") |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
149 |
fun replace_last_name [n] = [prefix ^ n ^ suffix] |
79490 | 150 |
| replace_last_name (n::ns) = n :: (replace_last_name ns) |
151 |
| replace_last_name _ = error "Internal error: Invalid function name to convert" |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
152 |
val parts = split_name name |
79490 | 153 |
in |
154 |
String.concatWith "." (replace_last_name parts) |
|
155 |
end |
|
81147 | 156 |
fun fun_name_to_time ctxt s name = fun_name_to_time' ctxt s false name |
79490 | 157 |
(* Count number of arguments of a function *) |
158 |
fun count_args (Type (n, [_,res])) = (if n = "fun" then 1 + count_args res else 0) |
|
159 |
| count_args _ = 0 |
|
160 |
(* Check if number of arguments matches function *) |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
161 |
fun check_args s (t, args) = |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
162 |
(if length args = count_args (type_of t) then () |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
163 |
else error ("Partial applications/Lambdas not allowed (" ^ s ^ ")")) |
79490 | 164 |
(* Removes Abs *) |
165 |
fun rem_abs f (Abs (_,_,t)) = rem_abs f t |
|
166 |
| rem_abs f t = f t |
|
167 |
(* Map right side of equation *) |
|
168 |
fun map_r f (pT $ (eq $ l $ r)) = (pT $ (eq $ l $ f r)) |
|
169 |
| map_r _ _ = error "Internal error: No right side of equation found" |
|
170 |
(* Get left side of equation *) |
|
171 |
fun get_l (_ $ (_ $ l $ _)) = l |
|
172 |
| get_l _ = error "Internal error: No left side of equation found" |
|
173 |
(* Get right side of equation *) |
|
174 |
fun get_r (_ $ (_ $ _ $ r)) = r |
|
175 |
| get_r _ = error "Internal error: No right side of equation found" |
|
176 |
(* Return name of Const *) |
|
177 |
fun Const_name (Const (nm,_)) = SOME nm |
|
178 |
| Const_name _ = NONE |
|
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
179 |
fun is_Used (Type ("Product_Type.prod", _)) = true |
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
180 |
| is_Used _ = false |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
181 |
(* Custom compare function for types ignoring variable names *) |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
182 |
fun typ_comp (Type (A,a)) (Type (B,b)) = (A = B) andalso List.foldl (fn ((c,i),s) => typ_comp c i andalso s) true (ListPair.zip (a, b)) |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
183 |
| typ_comp (Type _) _ = false |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
184 |
| typ_comp _ (Type _) = false |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
185 |
| typ_comp _ _ = true |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
186 |
fun const_comp (Const (nm,T)) (Const (nm',T')) = nm = nm' andalso typ_comp T T' |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
187 |
| const_comp _ _ = false |
79490 | 188 |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
189 |
fun time_term ctxt s (Const (nm,T)) = |
79490 | 190 |
let |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
191 |
val T_nm = fun_name_to_time ctxt s nm |
79490 | 192 |
val T_T = change_typ T |
193 |
in |
|
194 |
(SOME (Syntax.check_term ctxt (Const (T_nm,T_T)))) |
|
195 |
handle (ERROR _) => |
|
196 |
case Syntax.read_term ctxt (Long_Name.base_name T_nm) |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
197 |
of (Const (T_nm,T_T)) => |
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
198 |
let |
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
199 |
fun col_Used i (Type ("fun", [Type ("fun", _), Ts])) (Type ("fun", [T', Ts'])) = |
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
200 |
(if is_Used T' then [i] else []) @ col_Used (i+1) Ts Ts' |
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
201 |
| col_Used i (Type ("fun", [_, Ts])) (Type ("fun", [_, Ts'])) = col_Used (i+1) Ts Ts' |
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
202 |
| col_Used _ _ _ = [] |
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
203 |
in |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
204 |
SOME (Const (T_nm,change_typ' (contains (col_Used 0 T T_T)) T)) |
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
205 |
end |
79490 | 206 |
| _ => error ("Timing function of " ^ nm ^ " is not defined") |
207 |
end |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
208 |
| time_term _ _ _ = error "Internal error: No valid function given" |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
209 |
|
79490 | 210 |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
211 |
type 'a wctxt = { |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
212 |
ctxt: local_theory, |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
213 |
origins: term list, |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
214 |
f: term -> 'a |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
215 |
} |
79490 | 216 |
type 'a converter = { |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
217 |
constc : 'a wctxt -> term -> 'a, |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
218 |
funcc : 'a wctxt -> term -> term list -> 'a, |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
219 |
ifc : 'a wctxt -> typ -> term -> term -> term -> 'a, |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
220 |
casec : 'a wctxt -> term -> term list -> 'a, |
81147 | 221 |
letc : 'a wctxt -> typ -> term -> (string * typ) list -> term -> 'a |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
222 |
} |
79490 | 223 |
|
224 |
(* Walks over term and calls given converter *) |
|
81147 | 225 |
(* get rid and use Term.strip_abs.eta especially for lambdas *) |
81255 | 226 |
fun list_abs ([], t) = t |
227 |
| list_abs (a::abs,t) = list_abs (abs,t) |> absfree a |
|
79494
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
228 |
fun walk ctxt (origin: term list) (conv as {ifc, casec, funcc, letc, ...} : 'a converter) (t as _ $ _) = |
79490 | 229 |
let |
81147 | 230 |
val (f, args) = strip_comb t |
79490 | 231 |
val this = (walk ctxt origin conv) |
232 |
val _ = (case f of Abs _ => error "Lambdas not supported" | _ => ()) |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
233 |
val wctxt = {ctxt = ctxt, origins = origin, f = this} |
79490 | 234 |
in |
235 |
(if is_if f then |
|
236 |
(case f of (Const (_,T)) => |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
237 |
(case args of [cond, t, f] => ifc wctxt T cond t f |
79490 | 238 |
| _ => error "Partial applications not supported (if)") |
239 |
| _ => error "Internal error: invalid if term") |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
240 |
else if is_case f then casec wctxt f args |
79490 | 241 |
else if is_let f then |
242 |
(case f of (Const (_,lT)) => |
|
81147 | 243 |
(case args of [exp, t] => |
81255 | 244 |
let val (abs,t) = Term.strip_abs_eta 1 t in letc wctxt lT exp abs t end |
79490 | 245 |
| _ => error "Partial applications not allowed (let)") |
246 |
| _ => error "Internal error: invalid let term") |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
247 |
else funcc wctxt f args) |
79490 | 248 |
end |
249 |
| walk ctxt origin (conv as {constc, ...}) c = |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
250 |
constc {ctxt = ctxt, origins = origin, f = walk ctxt origin conv} c |
81147 | 251 |
fun Ifunc (wctxt: term wctxt) t args = list_comb (#f wctxt t,map (#f wctxt) args) |
252 |
val Iconst = K I |
|
253 |
fun Iif (wctxt: term wctxt) T cond tt tf = |
|
254 |
Const (@{const_name "HOL.If"}, T) $ (#f wctxt cond) $ (#f wctxt tt) $ (#f wctxt tf) |
|
255 |
fun Icase (wctxt: term wctxt) t cs = list_comb (#f wctxt t,map (#f wctxt) cs) |
|
256 |
fun Ilet (wctxt: term wctxt) lT exp abs t = |
|
81255 | 257 |
Const (@{const_name "HOL.Let"},lT) $ (#f wctxt exp) $ list_abs (abs, #f wctxt t) |
79490 | 258 |
|
259 |
(* 1. Fix all terms *) |
|
260 |
(* Exchange Var in types and terms to Free *) |
|
81147 | 261 |
fun freeTerms (Var(ixn,T)) = Free (fst ixn, T) |
262 |
| freeTerms t = t |
|
263 |
fun freeTypes (TVar ((t, _), T)) = TFree (t, T) |
|
264 |
| freeTypes t = t |
|
79542 | 265 |
|
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
266 |
fun noFun (Type ("fun", _)) = error "Functions in datatypes are not allowed in case constructions" |
80643
11b8f2e4c3d2
branches of case expressions may need to be eta-expanded
nipkow
parents:
80636
diff
changeset
|
267 |
| noFun T = T |
79490 | 268 |
fun casecBuildBounds n t = if n > 0 then casecBuildBounds (n-1) (t $ (Bound (n-1))) else t |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
269 |
fun casecAbs wctxt n (Type ("fun",[T,Tr])) (Abs (v,Ta,t)) = (map_atyps noFun T; Abs (v,Ta,casecAbs wctxt n Tr t)) |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
270 |
| casecAbs wctxt n (Type ("fun",[T,Tr])) t = |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
271 |
(map_atyps noFun T; Abs ("uu",T,casecAbs wctxt (n + 1) Tr t)) |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
272 |
| casecAbs wctxt n _ t = (#f wctxt) (casecBuildBounds n (Term.incr_bv n 0 t)) |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
273 |
fun fixCasecCases _ _ [t] = [t] |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
274 |
| fixCasecCases wctxt (Type (_,[T,Tr])) (t::ts) = casecAbs wctxt 0 T t :: fixCasecCases wctxt Tr ts |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
275 |
| fixCasecCases _ _ _ = error "Internal error: invalid case types/terms" |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
276 |
fun fixCasec wctxt (t as Const (_,T)) args = |
81147 | 277 |
(check_args "cases" (t,args); list_comb (t,fixCasecCases wctxt T args)) |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
278 |
| fixCasec _ _ _ = error "Internal error: invalid case term" |
79490 | 279 |
|
81147 | 280 |
fun shortFunc fixedNum (Const (nm,T)) = |
281 |
Const (nm,T |> strip_type |>> drop fixedNum |> (op --->)) |
|
282 |
| shortFunc _ _ = error "Internal error: Invalid term" |
|
283 |
fun shortApp fixedNum (c, args) = |
|
284 |
(shortFunc fixedNum c, drop fixedNum args) |
|
285 |
fun shortOriginFunc (term: term list) fixedNum (f as (c as Const (_,_), _)) = |
|
286 |
if contains' const_comp term c then shortApp fixedNum f else f |
|
287 |
| shortOriginFunc _ _ t = t |
|
81255 | 288 |
val _ = strip_abs |
289 |
fun map_abs f (t as Abs _) = t |> strip_abs ||> f |> list_abs |
|
290 |
| map_abs _ t = t |
|
81147 | 291 |
fun fixTerms ctxt (term: term list) (fixedNum: int) (t as pT $ (eq $ l $ r)) = |
79490 | 292 |
let |
81147 | 293 |
val _ = check_args "args" (strip_comb (get_l t)) |
294 |
val l' = shortApp fixedNum (strip_comb l) |> list_comb |
|
295 |
val shortOriginFunc' = shortOriginFunc (term |> map (fst o strip_comb)) fixedNum |
|
296 |
val r' = walk ctxt term { |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
297 |
funcc = (fn wctxt => fn t => fn args => |
81255 | 298 |
(check_args "func" (t,args); (#f wctxt t, map (#f wctxt) args) |> shortOriginFunc' |> list_comb)), |
299 |
constc = fn wctxt => map_abs (#f wctxt), |
|
81147 | 300 |
ifc = Iif, |
79490 | 301 |
casec = fixCasec, |
81147 | 302 |
letc = (fn wctxt => fn expT => fn exp => fn abs => fn t => |
81255 | 303 |
(Const (@{const_name "HOL.Let"},expT) $ (#f wctxt exp) $ list_abs (abs, #f wctxt t))) |
81147 | 304 |
} r |
305 |
in |
|
306 |
pT $ (eq $ l' $ r') |
|
79490 | 307 |
end |
81147 | 308 |
| fixTerms _ _ _ _ = error "Internal error: invalid term" |
79490 | 309 |
|
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
310 |
(* 2. Check for properties about the function *) |
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
311 |
(* 2.1 Check if function is recursive *) |
79490 | 312 |
fun or f (a,b) = f a orelse b |
313 |
fun find_rec ctxt term = (walk ctxt term { |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
314 |
funcc = (fn wctxt => fn t => fn args => |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
315 |
List.exists (fn term => (Const_name t) = (Const_name term)) term |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
316 |
orelse List.foldr (or (#f wctxt)) false args), |
81255 | 317 |
constc = fn wctxt => fn t => case t of |
318 |
Abs _ => t |> strip_abs |> snd |> (#f wctxt) |
|
319 |
| _ => false, |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
320 |
ifc = (fn wctxt => fn _ => fn cond => fn tt => fn tf => |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
321 |
(#f wctxt) cond orelse (#f wctxt) tt orelse (#f wctxt) tf), |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
322 |
casec = (fn wctxt => fn t => fn cs => |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
323 |
(#f wctxt) t orelse List.foldr (or (rem_abs (#f wctxt))) false cs), |
81147 | 324 |
letc = (fn wctxt => fn _ => fn exp => fn _ => fn t => |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
325 |
(#f wctxt) exp orelse (#f wctxt) t) |
79490 | 326 |
}) o get_r |
79494
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
327 |
fun is_rec ctxt (term: term list) = List.foldr (or (find_rec ctxt term)) false |
79490 | 328 |
|
329 |
(* 3. Convert equations *) |
|
330 |
(* Some Helper *) |
|
331 |
val plusTyp = @{typ "nat => nat => nat"} |
|
332 |
fun plus (SOME a) (SOME b) = SOME (Const (@{const_name "Groups.plus"}, plusTyp) $ a $ b) |
|
333 |
| plus (SOME a) NONE = SOME a |
|
334 |
| plus NONE (SOME b) = SOME b |
|
335 |
| plus NONE NONE = NONE |
|
336 |
fun opt_term NONE = HOLogic.zero |
|
337 |
| opt_term (SOME t) = t |
|
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
338 |
fun use_origin (Free (nm, T as Type ("fun",_))) = HOLogic.mk_fst (Free (nm,HOLogic.mk_prodT (T, change_typ T))) |
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
339 |
| use_origin t = t |
79490 | 340 |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
341 |
(* Conversion of function term *) |
81147 | 342 |
fun fun_to_time' ctxt (origin: term list) second (func as Const (nm,T)) = |
79490 | 343 |
let |
81147 | 344 |
val origin' = map (fst o strip_comb) origin |
79490 | 345 |
in |
81147 | 346 |
if contains' const_comp origin' func then SOME (Free (func |> Term.term_name |> fun_name_to_time' ctxt true second, change_typ T)) else |
79490 | 347 |
if Zero_Funcs.is_zero (Proof_Context.theory_of ctxt) (nm,T) then NONE else |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
348 |
time_term ctxt false func |
79490 | 349 |
end |
81147 | 350 |
| fun_to_time' _ _ _ (Free (nm,T)) = |
351 |
SOME (HOLogic.mk_snd (Free (nm,HOLogic.mk_prodT (T,change_typ T)))) |
|
352 |
| fun_to_time' _ _ _ _ = error "Internal error: invalid function to convert" |
|
353 |
fun fun_to_time context origin func = fun_to_time' context origin false func |
|
79490 | 354 |
|
355 |
(* Convert arguments of left side of a term *) |
|
81147 | 356 |
fun conv_arg _ (Free (nm,T as Type("fun",_))) = |
357 |
Free (nm, HOLogic.mk_prodT (T, change_typ' (K false) T)) |
|
358 |
| conv_arg _ x = x |
|
359 |
fun conv_args ctxt = map (conv_arg ctxt) |
|
79490 | 360 |
|
361 |
(* Handle function calls *) |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
362 |
fun build_zero (Type ("fun", [T, R])) = Abs ("uu", T, build_zero R) |
79490 | 363 |
| build_zero _ = zero |
81147 | 364 |
fun funcc_use_origin (Free (nm, T as Type ("fun",_))) = |
365 |
HOLogic.mk_fst (Free (nm,HOLogic.mk_prodT (T, change_typ T))) |
|
366 |
| funcc_use_origin t = t |
|
367 |
fun funcc_conv_arg _ _ (t as (_ $ _)) = map_aterms funcc_use_origin t |
|
368 |
| funcc_conv_arg _ u (Free (nm, T as Type ("fun",_))) = |
|
369 |
if u then Free (nm, HOLogic.mk_prodT (T, change_typ T)) |
|
370 |
else HOLogic.mk_snd (Free (nm,HOLogic.mk_prodT (T,change_typ T))) |
|
81255 | 371 |
| funcc_conv_arg wctxt true (f as Const (_,Type ("fun",_))) = |
372 |
HOLogic.mk_prod (f, funcc_conv_arg wctxt false f) |
|
81147 | 373 |
| funcc_conv_arg wctxt false (f as Const (_,T as Type ("fun",_))) = |
374 |
Option.getOpt (fun_to_time (#ctxt wctxt) (#origins wctxt) f, build_zero T) |
|
81255 | 375 |
| funcc_conv_arg wctxt false (f as Abs _) = |
376 |
f |
|
377 |
|> Term.strip_abs_eta ((length o fst o strip_type o type_of) f) |
|
378 |
||> #f wctxt ||> opt_term |
|
379 |
|> list_abs |
|
380 |
| funcc_conv_arg wctxt true (f as Abs _) = |
|
381 |
let |
|
382 |
val f' = f |
|
383 |
|> Term.strip_abs_eta ((length o fst o strip_type o type_of) f) |
|
384 |
||> map_aterms funcc_use_origin |
|
385 |
|> list_abs |
|
386 |
in |
|
387 |
HOLogic.mk_prod (f', funcc_conv_arg wctxt false f) |
|
388 |
end |
|
81147 | 389 |
| funcc_conv_arg _ _ t = t |
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
390 |
|
81147 | 391 |
fun funcc_conv_args _ _ [] = [] |
392 |
| funcc_conv_args wctxt (Type ("fun", [t, ts])) (a::args) = |
|
393 |
funcc_conv_arg wctxt (is_Used t) a :: funcc_conv_args wctxt ts args |
|
394 |
| funcc_conv_args _ _ _ = error "Internal error: Non matching type" |
|
395 |
fun funcc wctxt func args = |
|
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
396 |
let |
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
397 |
fun get_T (Free (_,T)) = T |
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
398 |
| get_T (Const (_,T)) = T |
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
399 |
| get_T (_ $ (Free (_,Type (_, [_, T])))) = T (* Case of snd was constructed *) |
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
400 |
| get_T _ = error "Internal error: Forgotten type" |
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
401 |
in |
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
402 |
List.foldr (I #-> plus) |
81255 | 403 |
(case fun_to_time (#ctxt wctxt) (#origins wctxt) func (* add case for abs *) |
404 |
of SOME t => SOME (list_comb (t, funcc_conv_args wctxt (get_T t) args)) |
|
405 |
| NONE => NONE) |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
406 |
(map (#f wctxt) args) |
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
407 |
end |
79490 | 408 |
|
409 |
(* Handle case terms *) |
|
410 |
fun casecIsCase (Type (n1, [_,Type (n2, _)])) = (n1 = "fun" andalso n2 = "fun") |
|
411 |
| casecIsCase _ = false |
|
412 |
fun casecLastTyp (Type (n, [T1,T2])) = Type (n, [T1, change_typ T2]) |
|
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
413 |
| casecLastTyp _ = error "Internal error: Invalid case type" |
79490 | 414 |
fun casecTyp (Type (n, [T1, T2])) = |
415 |
Type (n, [change_typ T1, (if casecIsCase T2 then casecTyp else casecLastTyp) T2]) |
|
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
416 |
| casecTyp _ = error "Internal error: Invalid case type" |
79490 | 417 |
fun casecAbs f (Abs (v,Ta,t)) = (case casecAbs f t of (nconst,t) => (nconst,Abs (v,Ta,t))) |
418 |
| casecAbs f t = (case f t of NONE => (false,HOLogic.zero) | SOME t => (true,t)) |
|
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
419 |
fun casecArgs _ [t] = (false, [map_aterms use_origin t]) |
79490 | 420 |
| casecArgs f (t::ar) = |
421 |
(case casecAbs f t of (nconst, tt) => |
|
422 |
casecArgs f ar ||> (fn ar => tt :: ar) |>> (if nconst then K true else I)) |
|
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
423 |
| casecArgs _ _ = error "Internal error: Invalid case term" |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
424 |
fun casec wctxt (Const (t,T)) args = |
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
425 |
if not (casecIsCase T) then error "Internal error: Invalid case type" else |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
426 |
let val (nconst, args') = casecArgs (#f wctxt) args in |
79490 | 427 |
plus |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
428 |
((#f wctxt) (List.last args)) |
79490 | 429 |
(if nconst then |
81147 | 430 |
SOME (list_comb (Const (t,casecTyp T), args')) |
79490 | 431 |
else NONE) |
432 |
end |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
433 |
| casec _ _ _ = error "Internal error: Invalid case term" |
79490 | 434 |
|
435 |
(* Handle if terms -> drop the term if true and false terms are zero *) |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
436 |
fun ifc wctxt _ cond tt ft = |
79490 | 437 |
let |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
438 |
val f = #f wctxt |
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
439 |
val rcond = map_aterms use_origin cond |
79490 | 440 |
val tt = f tt |
441 |
val ft = f ft |
|
442 |
in |
|
79969 | 443 |
plus (f cond) (case (tt,ft) of (NONE, NONE) => NONE | _ => |
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
444 |
if tt = ft then tt else |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
445 |
(SOME ((Const (@{const_name "HOL.If"}, @{typ "bool \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"})) $ rcond $ (opt_term tt) $ (opt_term ft)))) |
79490 | 446 |
end |
447 |
||
81255 | 448 |
fun letc_lambda wctxt T (t as Abs _) = |
449 |
HOLogic.mk_prod (map_aterms use_origin t, |
|
450 |
Term.strip_abs_eta (strip_type T |> fst |> length) t ||> #f wctxt ||> opt_term |> list_abs) |
|
451 |
| letc_lambda _ _ t = map_aterms use_origin t |
|
452 |
fun letc wctxt expT exp ([(nm,_)]) t = |
|
453 |
plus (#f wctxt exp) |
|
454 |
(case #f wctxt t of SOME t' => |
|
455 |
(if Term.used_free nm t' |
|
456 |
then |
|
457 |
let |
|
458 |
val exp' = letc_lambda wctxt expT exp |
|
459 |
val t' = list_abs ([(nm,type_of exp')], t') |
|
460 |
in |
|
461 |
Const (@{const_name "HOL.Let"}, [type_of exp', type_of t'] ---> HOLogic.natT) $ exp' $ t' |
|
462 |
end |
|
463 |
else t') |> SOME |
|
464 |
| NONE => NONE) |
|
465 |
| letc _ _ _ _ _ = error "Unknown let state" |
|
466 |
||
467 |
fun constc _ (Const ("HOL.undefined", _)) = SOME (Const ("HOL.undefined", @{typ "nat"})) |
|
468 |
| constc _ _ = NONE |
|
79490 | 469 |
|
470 |
(* The converter for timing functions given to the walker *) |
|
81147 | 471 |
val converter : term option converter = { |
81255 | 472 |
constc = constc, |
81147 | 473 |
funcc = funcc, |
79490 | 474 |
ifc = ifc, |
475 |
casec = casec, |
|
476 |
letc = letc |
|
477 |
} |
|
79665 | 478 |
fun top_converter is_rec _ _ = opt_term o (fn exp => plus exp (if is_rec then SOME one else NONE)) |
79490 | 479 |
|
480 |
(* Use converter to convert right side of a term *) |
|
81147 | 481 |
fun to_time ctxt origin is_rec term = |
482 |
top_converter is_rec ctxt origin (walk ctxt origin converter term) |
|
79490 | 483 |
|
484 |
(* Converts a term to its running time version *) |
|
81147 | 485 |
fun convert_term ctxt (origin: term list) is_rec (pT $ (Const (eqN, _) $ l $ r)) = |
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
486 |
let |
81147 | 487 |
val (l_const, l_params) = strip_comb l |
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
488 |
in |
81147 | 489 |
pT |
490 |
$ (Const (eqN, @{typ "nat \<Rightarrow> nat \<Rightarrow> bool"}) |
|
491 |
$ (list_comb (l_const |> fun_to_time ctxt origin |> Option.valOf, l_params |> conv_args ctxt)) |
|
492 |
$ (to_time ctxt origin is_rec r)) |
|
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
493 |
end |
81147 | 494 |
| convert_term _ _ _ _ = error "Internal error: invalid term to convert" |
495 |
||
496 |
(* 3.5 Support for locales *) |
|
497 |
fun replaceFstSndFree ctxt (origin: term list) (rfst: term -> term) (rsnd: term -> term) = |
|
498 |
(walk ctxt origin { |
|
499 |
funcc = fn wctxt => fn t => fn args => |
|
500 |
case args of |
|
501 |
(f as Free _)::args => |
|
502 |
(case t of |
|
503 |
Const ("Product_Type.prod.fst", _) => |
|
504 |
list_comb (rfst (t $ f), map (#f wctxt) args) |
|
505 |
| Const ("Product_Type.prod.snd", _) => |
|
506 |
list_comb (rsnd (t $ f), map (#f wctxt) args) |
|
507 |
| t => list_comb (t, map (#f wctxt) (f :: args))) |
|
508 |
| args => list_comb (t, map (#f wctxt) args), |
|
509 |
constc = Iconst, |
|
510 |
ifc = Iif, |
|
511 |
casec = Icase, |
|
512 |
letc = Ilet |
|
513 |
}) |
|
79490 | 514 |
|
515 |
(* 4. Tactic to prove "f_dom n" *) |
|
516 |
fun time_dom_tac ctxt induct_rule domintros = |
|
517 |
(Induction.induction_tac ctxt true [] [[]] [] (SOME [induct_rule]) [] |
|
518 |
THEN_ALL_NEW ((K (auto_tac ctxt)) THEN' (fn i => FIRST' ( |
|
519 |
(if i <= length domintros then [Metis_Tactic.metis_tac [] ATP_Problem_Generate.combsN ctxt [List.nth (domintros, i-1)]] else []) @ |
|
520 |
[Metis_Tactic.metis_tac [] ATP_Problem_Generate.combsN ctxt domintros]) i))) |
|
521 |
||
522 |
||
81255 | 523 |
fun fix_definition (Const ("Pure.eq", _) $ l $ r) = HOLogic.mk_Trueprop (HOLogic.mk_eq (l,r)) |
524 |
| fix_definition t = t |
|
525 |
fun check_definition [t] = [t] |
|
526 |
| check_definition _ = error "Only a single definition is allowed" |
|
79490 | 527 |
fun get_terms theory (term: term) = |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
528 |
let |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
529 |
val equations = Spec_Rules.retrieve theory term |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
530 |
|> map #rules |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
531 |
|> map (map Thm.prop_of) |
79490 | 532 |
handle Empty => error "Function or terms of function not found" |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
533 |
in |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
534 |
equations |
81255 | 535 |
|> map (map fix_definition) |
81147 | 536 |
|> filter (List.exists |
537 |
(fn t => typ_comp (t |> get_l |> strip_comb |> fst |> dest_Const |> snd) (term |> strip_comb |> fst |> dest_Const |> snd))) |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
538 |
|> hd |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
539 |
end |
79490 | 540 |
|
81147 | 541 |
(* 5. Check for higher-order function if original function is used \<rightarrow> find simplifications *) |
542 |
fun find_used' T_t = |
|
543 |
let |
|
544 |
val (T_ident, T_args) = strip_comb (get_l T_t) |
|
545 |
||
546 |
fun filter_passed [] = [] |
|
547 |
| filter_passed ((f as Free (_, Type ("Product_Type.prod",[Type ("fun",_), Type ("fun", _)])))::args) = |
|
548 |
f :: filter_passed args |
|
549 |
| filter_passed (_::args) = filter_passed args |
|
550 |
val frees = (walk @{context} [] { |
|
551 |
funcc = (fn wctxt => fn t => fn args => |
|
552 |
(case t of (Const ("Product_Type.prod.snd", _)) => [] |
|
553 |
| _ => (if t = T_ident then [] else filter_passed args) |
|
554 |
@ List.foldr (fn (l,r) => (#f wctxt) l @ r) [] args)), |
|
555 |
constc = (K o K) [], |
|
556 |
ifc = (fn wctxt => fn _ => fn cond => fn tt => fn tf => (#f wctxt) cond @ (#f wctxt) tt @ (#f wctxt) tf), |
|
557 |
casec = (fn wctxt => fn _ => fn cs => List.foldr (fn (l,r) => (#f wctxt) l @ r) [] cs), |
|
558 |
letc = (fn wctxt => fn _ => fn exp => fn _ => fn t => (#f wctxt) exp @ (#f wctxt) t) |
|
559 |
}) (get_r T_t) |
|
560 |
fun build _ [] = [] |
|
561 |
| build i (a::args) = |
|
562 |
(if contains frees a then [(T_ident,i)] else []) @ build (i+1) args |
|
563 |
in |
|
564 |
build 0 T_args |
|
565 |
end |
|
566 |
fun find_simplifyble ctxt term terms = |
|
567 |
let |
|
568 |
val used = |
|
569 |
terms |
|
570 |
|> List.map find_used' |
|
571 |
|> List.foldr (op @) [] |
|
572 |
val change = |
|
573 |
Option.valOf o fun_to_time ctxt term |
|
574 |
fun detect t i (Type ("fun",_)::args) = |
|
575 |
(if contains used (change t,i) then [] else [i]) @ detect t (i+1) args |
|
576 |
| detect t i (_::args) = detect t (i+1) args |
|
577 |
| detect _ _ [] = [] |
|
578 |
in |
|
579 |
map (fn t => t |> type_of |> strip_type |> fst |> detect t 0) term |
|
580 |
end |
|
581 |
||
582 |
fun define_simp' term simplifyable ctxt = |
|
583 |
let |
|
584 |
val base_name = case Named_Target.locale_of ctxt of |
|
585 |
NONE => ctxt |> Proof_Context.theory_of |> Context.theory_base_name |
|
586 |
| SOME nm => nm |
|
587 |
||
588 |
val orig_name = term |> dest_Const_name |> split_name |> List.last |
|
589 |
val red_name = fun_name_to_time ctxt false orig_name |
|
590 |
val name = fun_name_to_time' ctxt true true orig_name |
|
591 |
val full_name = base_name ^ "." ^ name |
|
592 |
val def_name = red_name ^ "_def" |
|
593 |
val def = Binding.name def_name |
|
594 |
||
595 |
val canon = Syntax.read_term (Local_Theory.exit ctxt) name |> strip_comb |
|
596 |
val canonFrees = canon |> snd |
|
597 |
val canonType = canon |> fst |> dest_Const_type |> strip_type |> fst |> take (length canonFrees) |
|
598 |
||
599 |
val types = term |> dest_Const_type |> strip_type |> fst |
|
600 |
val vars = Variable.variant_fixes (map (K "") types) ctxt |> fst |
|
601 |
fun l_typs' i ((T as (Type ("fun",_)))::types) = |
|
602 |
(if contains simplifyable i |
|
603 |
then change_typ T |
|
604 |
else HOLogic.mk_prodT (T,change_typ T)) |
|
605 |
:: l_typs' (i+1) types |
|
606 |
| l_typs' i (T::types) = T :: l_typs' (i+1) types |
|
607 |
| l_typs' _ [] = [] |
|
608 |
val l_typs = l_typs' 0 types |
|
609 |
val lhs = |
|
610 |
List.foldl (fn ((v,T),t) => t $ Free (v,T)) (Free (red_name,l_typs ---> HOLogic.natT)) (ListPair.zip (vars,l_typs)) |
|
611 |
fun fixType (TFree _) = HOLogic.natT |
|
612 |
| fixType T = T |
|
613 |
fun fixUnspecified T = T |> strip_type ||> fixType |> (op --->) |
|
614 |
fun r_terms' i (v::vars) ((T as (Type ("fun",_)))::types) = |
|
615 |
(if contains simplifyable i |
|
616 |
then HOLogic.mk_prod (Const ("HOL.undefined", fixUnspecified T), Free (v,change_typ T)) |
|
617 |
else Free (v,HOLogic.mk_prodT (T,change_typ T))) |
|
618 |
:: r_terms' (i+1) vars types |
|
619 |
| r_terms' i (v::vars) (T::types) = Free (v,T) :: r_terms' (i+1) vars types |
|
620 |
| r_terms' _ _ _ = [] |
|
621 |
val r_terms = r_terms' 0 vars types |
|
622 |
val full_type = (r_terms |> map (type_of) ---> HOLogic.natT) |
|
623 |
val full = list_comb (Const (full_name,canonType ---> full_type), canonFrees) |
|
624 |
val rhs = list_comb (full, r_terms) |
|
625 |
val eq = (lhs, rhs) |> HOLogic.mk_eq |> HOLogic.mk_Trueprop |
|
626 |
val _ = Pretty.writeln (Pretty.block [Pretty.str "Defining simplified version:\n", |
|
627 |
Syntax.pretty_term ctxt eq]) |
|
628 |
||
629 |
val (_, ctxt') = Specification.definition NONE [] [] ((def, []), eq) ctxt |
|
630 |
||
631 |
in |
|
632 |
((def_name, orig_name), ctxt') |
|
633 |
end |
|
634 |
fun define_simp simpables ctxt = |
|
635 |
let |
|
636 |
fun cond ((term,simplifyable),(defs,ctxt)) = |
|
637 |
define_simp' term simplifyable ctxt |>> (fn def => def :: defs) |
|
638 |
in |
|
639 |
List.foldr cond ([], ctxt) simpables |
|
640 |
end |
|
641 |
||
642 |
||
643 |
fun replace from to = |
|
644 |
map (map_aterms (fn t => if t = from then to else t)) |
|
645 |
fun replaceAll [] = I |
|
646 |
| replaceAll ((from,to)::xs) = replaceAll xs o replace from to |
|
647 |
fun calculateSimplifications ctxt T_terms term simpables = |
|
648 |
let |
|
649 |
(* Show where a simplification can take place *) |
|
650 |
fun reportReductions (t,(i::is)) = |
|
651 |
(Pretty.writeln (Pretty.str |
|
652 |
((Term.term_name t |> fun_name_to_time ctxt true) |
|
653 |
^ " can be simplified because only the time-function component of parameter " |
|
654 |
^ (Int.toString (i + 1)) ^ " is used. ")); |
|
655 |
reportReductions (t,is)) |
|
656 |
| reportReductions (_,[]) = () |
|
657 |
val _ = simpables |
|
658 |
|> map reportReductions |
|
659 |
||
660 |
(* Register definitions for simplified function *) |
|
661 |
val (reds, ctxt) = define_simp simpables ctxt |
|
662 |
||
663 |
fun genRetype (Const (nm,T),is) = |
|
664 |
let |
|
665 |
val T_name = fun_name_to_time ctxt true nm |> split_name |> List.last |
|
666 |
val from = Free (T_name,change_typ T) |
|
667 |
val to = Free (T_name,change_typ' (not o contains is) T) |
|
668 |
in |
|
669 |
(from,to) |
|
670 |
end |
|
671 |
| genRetype _ = error "Internal error: invalid term" |
|
672 |
val retyping = map genRetype simpables |
|
673 |
||
674 |
fun replaceArgs (pT $ (eq $ l $ r)) = |
|
675 |
let |
|
676 |
val (t,params) = strip_comb l |
|
677 |
fun match (Const (f_nm,_),_) = |
|
678 |
(fun_name_to_time ctxt true f_nm |> Long_Name.base_name) = (dest_Free t |> fst) |
|
679 |
| match _ = false |
|
680 |
val simps = List.find match simpables |> Option.valOf |> snd |
|
681 |
||
682 |
fun dest_Prod_snd (Free (nm, Type (_, [_, T2]))) = |
|
683 |
Free (fun_name_to_time ctxt false nm, T2) |
|
684 |
| dest_Prod_snd _ = error "Internal error: Argument is not a pair" |
|
685 |
fun rep _ [] = ([],[]) |
|
686 |
| rep i (x::xs) = |
|
687 |
let |
|
688 |
val (rs,args) = rep (i+1) xs |
|
689 |
in |
|
690 |
if contains simps i |
|
691 |
then (x::rs,dest_Prod_snd x::args) |
|
692 |
else (rs,x::args) |
|
693 |
end |
|
694 |
val (rs,params) = rep 0 params |
|
695 |
fun fFst _ = error "Internal error: Invalid term to simplify" |
|
696 |
fun fSnd (t as (Const _ $ f)) = |
|
697 |
(if contains rs f |
|
698 |
then dest_Prod_snd f |
|
699 |
else t) |
|
700 |
| fSnd t = t |
|
701 |
in |
|
702 |
(pT $ (eq |
|
703 |
$ (list_comb (t,params)) |
|
704 |
$ (replaceFstSndFree ctxt term fFst fSnd r |
|
705 |
|> (fn t => replaceAll (map (fn t => (t,dest_Prod_snd t)) rs) [t]) |
|
706 |
|> hd |
|
707 |
) |
|
708 |
)) |
|
709 |
end |
|
710 |
| replaceArgs _ = error "Internal error: Invalid term" |
|
711 |
||
712 |
(* Calculate reduced terms *) |
|
713 |
val T_terms_red = T_terms |
|
714 |
|> replaceAll retyping |
|
715 |
|> map replaceArgs |
|
716 |
||
717 |
val _ = print_lemma ctxt reds T_terms_red |
|
718 |
val _ = |
|
719 |
Pretty.writeln (Pretty.str "If you do not want the simplified T function, use \"time_fun [no_simp]\"") |
|
720 |
in |
|
721 |
ctxt |
|
722 |
end |
|
723 |
||
79490 | 724 |
(* Register timing function of a given function *) |
81147 | 725 |
fun reg_time_func (lthy: local_theory) (term: term list) (terms: term list) print simp = |
79490 | 726 |
let |
727 |
val _ = |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
728 |
case time_term lthy true (hd term) |
79490 | 729 |
handle (ERROR _) => NONE |
79494
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
730 |
of SOME _ => error ("Timing function already declared: " ^ (Term.term_name (hd term))) |
79490 | 731 |
| NONE => () |
732 |
||
81147 | 733 |
(* Number of terms fixed by locale *) |
734 |
val fixedNum = term |
|
735 |
|> hd |
|
736 |
|> strip_comb |> snd |
|
737 |
|> length |
|
738 |
||
79490 | 739 |
(* 1. Fix all terms *) |
740 |
(* Exchange Var in types and terms to Free and check constraints *) |
|
79969 | 741 |
val terms = map |
81147 | 742 |
(map_aterms freeTerms |
743 |
#> map_types (map_atyps freeTypes) |
|
744 |
#> fixTerms lthy term fixedNum) |
|
79542 | 745 |
terms |
81147 | 746 |
val fixedFrees = (hd term) |> strip_comb |> snd |> take fixedNum |
747 |
val fixedFreesNames = map (fst o dest_Free) fixedFrees |
|
748 |
val term = map (shortFunc fixedNum o fst o strip_comb) term |
|
749 |
||
81255 | 750 |
fun correctTerm term = |
751 |
let |
|
752 |
val get_f = fst o strip_comb o get_l |
|
753 |
in |
|
754 |
List.find (fn t => (dest_Const_name o get_f) t = dest_Const_name term) terms |
|
755 |
|> Option.valOf |> get_f |
|
756 |
end |
|
757 |
val term = map correctTerm term |
|
79490 | 758 |
|
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
759 |
(* 2. Find properties about the function *) |
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
760 |
(* 2.1 Check if function is recursive *) |
79490 | 761 |
val is_rec = is_rec lthy term terms |
762 |
||
763 |
(* 3. Convert every equation |
|
764 |
- Change type of toplevel equation from _ \<Rightarrow> _ \<Rightarrow> bool to nat \<Rightarrow> nat \<Rightarrow> bool |
|
765 |
- On left side change name of function to timing function |
|
766 |
- Convert right side of equation with conversion schema |
|
767 |
*) |
|
81147 | 768 |
fun fFst (t as (Const (_,T) $ Free (nm,_))) = |
769 |
(if contains fixedFreesNames nm |
|
770 |
then Free (nm,strip_type T |>> tl |> (op --->)) |
|
771 |
else t) |
|
772 |
| fFst t = t |
|
773 |
fun fSnd (t as (Const (_,T) $ Free (nm,_))) = |
|
774 |
(if contains fixedFreesNames nm |
|
775 |
then Free (fun_name_to_time lthy false nm,strip_type T |>> tl |> (op --->)) |
|
776 |
else t) |
|
777 |
| fSnd t = t |
|
778 |
val T_terms = map (convert_term lthy term is_rec) terms |
|
779 |
|> map (map_r (replaceFstSndFree lthy term fFst fSnd)) |
|
780 |
||
781 |
val simpables = (if simp |
|
782 |
then find_simplifyble lthy term T_terms |
|
783 |
else map (K []) term) |
|
784 |
|> (fn s => ListPair.zip (term,s)) |
|
785 |
(* Determine if something is simpable, if so rename everything *) |
|
786 |
val simpable = simpables |> map snd |> exists (not o null) |
|
787 |
(* Rename to secondary if simpable *) |
|
788 |
fun genRename (t,_) = |
|
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
789 |
let |
81147 | 790 |
val old = fun_to_time lthy term t |> Option.valOf |
791 |
val new = fun_to_time' lthy term true t |> Option.valOf |
|
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
792 |
in |
81147 | 793 |
(old,new) |
80624
9f8034d29365
time_function T_map can now be generated automatically.
nipkow
parents:
80036
diff
changeset
|
794 |
end |
81147 | 795 |
val can_T_terms = if simpable |
796 |
then replaceAll (map genRename simpables) T_terms |
|
797 |
else T_terms |
|
79490 | 798 |
|
81147 | 799 |
(* 4. Register function and prove completeness *) |
79494
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
800 |
val names = map Term.term_name term |
81147 | 801 |
val timing_names = map (fun_name_to_time' lthy true simpable) names |
79494
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
802 |
val bindings = map (fn nm => (Binding.name nm, NONE, NoSyn)) timing_names |
79490 | 803 |
fun pat_completeness_auto ctxt = |
804 |
Pat_Completeness.pat_completeness_tac ctxt 1 THEN auto_tac ctxt |
|
81147 | 805 |
val specs = map (fn eq => (((Binding.empty, []), eq), [], [])) can_T_terms |
79490 | 806 |
|
81147 | 807 |
(* Context for printing without showing question marks *) |
808 |
val print_ctxt = lthy |
|
809 |
|> Config.put show_question_marks false |
|
810 |
|> Config.put show_sorts false (* Change it for debugging *) |
|
811 |
val print_ctxt = List.foldl (fn (term, ctxt) => Variable.add_fixes_implicit term ctxt) print_ctxt (term @ T_terms) |
|
812 |
(* Print result if print *) |
|
813 |
val _ = if not print then () else |
|
814 |
let |
|
815 |
val nms = map (dest_Const_name) term |
|
816 |
val typs = map (dest_Const_type) term |
|
817 |
in |
|
818 |
print_timing' print_ctxt { names=nms, terms=terms, typs=typs } |
|
819 |
{ names=timing_names, terms=can_T_terms, typs=map change_typ typs } |
|
820 |
end |
|
821 |
||
79490 | 822 |
(* For partial functions sequential=true is needed in order to support them |
823 |
We need sequential=false to support the automatic proof of termination over dom |
|
824 |
*) |
|
825 |
fun register seq = |
|
826 |
let |
|
827 |
val _ = (if seq then warning "Falling back on sequential function..." else ()) |
|
828 |
val fun_config = Function_Common.FunctionConfig |
|
79494
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
829 |
{sequential=seq, default=NONE, domintros=true, partials=false} |
79490 | 830 |
in |
831 |
Function.add_function bindings specs fun_config pat_completeness_auto lthy |
|
832 |
end |
|
833 |
||
81147 | 834 |
val (info,ctxt) = |
835 |
register false |
|
836 |
handle (ERROR _) => |
|
837 |
register true |
|
838 |
| Match => |
|
839 |
register true |
|
79490 | 840 |
|
81147 | 841 |
val ctxt = if simpable then calculateSimplifications ctxt T_terms term simpables else ctxt |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
842 |
in |
81147 | 843 |
(info,ctxt) |
79490 | 844 |
end |
81147 | 845 |
fun proove_termination (term: term list) terms (T_info: Function.info, lthy: local_theory) = |
79490 | 846 |
let |
847 |
(* Start proving the termination *) |
|
79494
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
848 |
val infos = SOME (map (Function.get_info lthy) term) handle Empty => NONE |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
849 |
val timing_names = map (fun_name_to_time lthy true o Term.term_name) term |
79490 | 850 |
|
851 |
(* Proof by lexicographic_order_tac *) |
|
852 |
val (time_info, lthy') = |
|
853 |
(Function.prove_termination NONE |
|
854 |
(Lexicographic_Order.lexicographic_order_tac false lthy) lthy) |
|
855 |
handle (ERROR _) => |
|
856 |
let |
|
857 |
val _ = warning "Falling back on proof over dom..." |
|
79494
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
858 |
val _ = (if length term > 1 then error "Proof over dom not supported for mutual recursive functions" else ()) |
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
859 |
|
79490 | 860 |
fun args (a$(Var ((nm,_),T))) = args a |> (fn ar => (nm,T)::ar) |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
861 |
| args (a$(Const (_,T))) = args a |> (fn ar => ("uu",T)::ar) |
79490 | 862 |
| args _ = [] |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
863 |
val dom_vars = |
81147 | 864 |
terms |> hd |> get_l |> map_types (map_atyps freeTypes) |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
865 |
|> args |> Variable.variant_frees lthy [] |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
866 |
val dom_args = |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
867 |
List.foldl (fn (t,p) => HOLogic.mk_prod ((Free t),p)) (Free (hd dom_vars)) (tl dom_vars) |
79490 | 868 |
|
79494
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
869 |
val {inducts, ...} = case infos of SOME [i] => i | _ => error "Proof over dom failed as no induct rule was found" |
79490 | 870 |
val induct = (Option.valOf inducts |> hd) |
871 |
||
79494
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
872 |
val domintros = Proof_Context.get_fact lthy (Facts.named (hd timing_names ^ ".domintros")) |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
873 |
val prop = HOLogic.mk_Trueprop (#dom T_info $ dom_args) |
79490 | 874 |
|
875 |
(* Prove a helper lemma *) |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
876 |
val dom_lemma = Goal.prove lthy (map fst dom_vars) [] prop |
79490 | 877 |
(fn {context, ...} => HEADGOAL (time_dom_tac context induct domintros)) |
878 |
(* Add dom_lemma to simplification set *) |
|
879 |
val simp_lthy = Simplifier.add_simp dom_lemma lthy |
|
880 |
in |
|
881 |
(* Use lemma to prove termination *) |
|
882 |
Function.prove_termination NONE |
|
883 |
(auto_tac simp_lthy) lthy |
|
884 |
end |
|
885 |
in |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
886 |
(time_info, lthy') |
79490 | 887 |
end |
81147 | 888 |
fun reg_and_proove_time_func (lthy: local_theory) (term: term list) (terms: term list) print simp = |
889 |
reg_time_func lthy term terms print simp |
|
890 |
|> proove_termination term terms |
|
79490 | 891 |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
892 |
fun isTypeClass' (Const (nm,_)) = |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
893 |
(case split_name nm |> rev |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
894 |
of (_::nm::_) => String.isSuffix "_class" nm |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
895 |
| _ => false) |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
896 |
| isTypeClass' _ = false |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
897 |
val isTypeClass = |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
898 |
(List.foldr (fn (a,b) => a orelse b) false) o (map isTypeClass') |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
899 |
|
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
900 |
fun detect_typ (ctxt: local_theory) (term: term) = |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
901 |
let |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
902 |
val class_term = (case term of Const (nm,_) => Syntax.read_term ctxt nm |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
903 |
| _ => error "Could not find term of class") |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
904 |
fun find_free (Type (_,class)) (Type (_,inst)) = |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
905 |
List.foldl (fn ((c,i),s) => (case s of NONE => find_free c i | t => t)) (NONE) (ListPair.zip (class, inst)) |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
906 |
| find_free (TFree _) (TFree _) = NONE |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
907 |
| find_free (TFree _) (Type (nm,_)) = SOME nm |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
908 |
| find_free _ _ = error "Unhandled case in detecting type" |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
909 |
in |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
910 |
find_free (type_of class_term) (type_of term) |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
911 |
|> Option.map (hd o rev o split_name) |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
912 |
end |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
913 |
|
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
914 |
fun set_suffix (fterms: term list) ctxt = |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
915 |
let |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
916 |
val isTypeClass = isTypeClass fterms |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
917 |
val _ = (if length fterms > 1 andalso isTypeClass then error "No mutual recursion inside instantiation allowed" else ()) |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
918 |
val suffix = (if isTypeClass then detect_typ ctxt (hd fterms) else NONE) |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
919 |
in |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
920 |
(case suffix of NONE => I | SOME s => Config.put bsuffix ("_" ^ s)) ctxt |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
921 |
end |
79969 | 922 |
|
81147 | 923 |
fun check_opts [] = false |
924 |
| check_opts ["no_simp"] = true |
|
925 |
| check_opts (a::_) = error ("Option " ^ a ^ " is not defined") |
|
926 |
||
79490 | 927 |
(* Convert function into its timing function (called by command) *) |
81147 | 928 |
fun reg_time_fun_cmd ((opts, funcs), thms) (ctxt: local_theory) = |
79490 | 929 |
let |
81147 | 930 |
val no_simp = check_opts opts |
79494
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
931 |
val fterms = map (Syntax.read_term ctxt) funcs |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
932 |
val ctxt = set_suffix fterms ctxt |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
933 |
val (_, ctxt') = reg_and_proove_time_func ctxt fterms |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
934 |
(case thms of NONE => get_terms ctxt (hd fterms) |
79490 | 935 |
| SOME thms => thms |> Attrib.eval_thms ctxt |> List.map Thm.prop_of) |
81147 | 936 |
true (not no_simp) |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
937 |
in ctxt' |
79490 | 938 |
end |
939 |
||
940 |
(* Convert function into its timing function (called by command) with termination proof provided by user*) |
|
81147 | 941 |
fun reg_time_function_cmd ((opts, funcs), thms) (ctxt: local_theory) = |
79490 | 942 |
let |
81147 | 943 |
val no_simp = check_opts opts |
79494
c7536609bb9b
translation to time functions now with canonical let.
nipkow
parents:
79490
diff
changeset
|
944 |
val fterms = map (Syntax.read_term ctxt) funcs |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
945 |
val ctxt = set_suffix fterms ctxt |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
946 |
val ctxt' = reg_time_func ctxt fterms |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
947 |
(case thms of NONE => get_terms ctxt (hd fterms) |
79490 | 948 |
| SOME thms => thms |> Attrib.eval_thms ctxt |> List.map Thm.prop_of) |
81147 | 949 |
true (not no_simp) |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
950 |
|> snd |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
951 |
in ctxt' |
79490 | 952 |
end |
953 |
||
79969 | 954 |
(* Convert function into its timing function (called by command) *) |
81147 | 955 |
fun reg_time_definition_cmd ((opts, funcs), thms) (ctxt: local_theory) = |
79969 | 956 |
let |
81147 | 957 |
val no_simp = check_opts opts |
79969 | 958 |
val fterms = map (Syntax.read_term ctxt) funcs |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
959 |
val ctxt = set_suffix fterms ctxt |
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
960 |
val (_, ctxt') = reg_and_proove_time_func ctxt fterms |
81255 | 961 |
(case thms of NONE => get_terms ctxt (hd fterms) |> check_definition |
79969 | 962 |
| SOME thms => thms |> Attrib.eval_thms ctxt |> List.map Thm.prop_of) |
81147 | 963 |
true (not no_simp) |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
964 |
in ctxt' |
79969 | 965 |
end |
966 |
||
81147 | 967 |
val parser = (Parse.opt_attribs >> map (fst o Token.name_of_src)) |
968 |
-- Scan.repeat1 Parse.prop |
|
969 |
-- Scan.option (Parse.keyword_improper "equations" -- Parse.thms1 >> snd) |
|
970 |
val _ = Toplevel.local_theory |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
971 |
val _ = Outer_Syntax.local_theory @{command_keyword "time_fun"} |
79490 | 972 |
"Defines runtime function of a function" |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
973 |
(parser >> reg_time_fun_cmd) |
79490 | 974 |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
975 |
val _ = Outer_Syntax.local_theory @{command_keyword "time_function"} |
79490 | 976 |
"Defines runtime function of a function" |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
977 |
(parser >> reg_time_function_cmd) |
79969 | 978 |
|
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
979 |
val _ = Outer_Syntax.local_theory @{command_keyword "time_definition"} |
79969 | 980 |
"Defines runtime function of a definition" |
80734
7054a1bc8347
new version of time_fun that works for classes; define T_length automatically now
nipkow
parents:
80657
diff
changeset
|
981 |
(parser >> reg_time_definition_cmd) |
79490 | 982 |
|
983 |
end |