266 |> forall_elim_list rng |
266 |> forall_elim_list rng |
267 end |
267 end |
268 |
268 |
269 val make_name = String.translate (fn #"." => "dot" | c => Char.toString c) |
269 val make_name = String.translate (fn #"." => "dot" | c => Char.toString c) |
270 |
270 |
271 fun make_free (x, ty) = Free (make_name x, ty) |
271 fun make_free (x, ty) = Thm.free (make_name x, ty); |
272 |
272 |
273 fun make_tfree a = |
273 fun make_tfree thy a = |
274 let val b = "'" ^ String.translate (fn #"?" => "t" | c => Char.toString c) a |
274 let val b = "'" ^ String.translate (fn #"?" => "t" | c => Char.toString c) a |
275 in TFree (b, \<^sort>\<open>type\<close>) end |
275 in Thm.global_ctyp_of thy (TFree (b, \<^sort>\<open>type\<close>)) end |
276 |
276 |
277 fun make_type thy (c, args) = |
277 fun make_type thy (c, args) = |
278 let |
278 let |
279 val d = |
279 val d = |
280 (case Import_Data.get_typ_map thy c of |
280 (case Import_Data.get_typ_map thy c of |
281 SOME d => d |
281 SOME d => d |
282 | NONE => Sign.full_bname thy (make_name c)) |
282 | NONE => Sign.full_bname thy (make_name c)) |
283 in Type (d, args) end |
283 val T = Thm.global_ctyp_of thy (Type (d, replicate (length args) dummyT)) |
|
284 in Thm.make_ctyp T args end |
284 |
285 |
285 fun make_const thy (c, ty) = |
286 fun make_const thy (c, ty) = |
286 let |
287 let |
287 val d = |
288 val d = |
288 (case Import_Data.get_const_map thy c of |
289 (case Import_Data.get_const_map thy c of |
289 SOME d => d |
290 SOME d => d |
290 | NONE => Sign.full_bname thy (make_name c)) |
291 | NONE => Sign.full_bname thy (make_name c)) |
291 in Const (d, ty) end |
292 in Thm.global_cterm_of thy (Const (d, Thm.typ_of ty)) end |
292 |
293 |
293 |
294 |
294 datatype state = |
295 datatype state = |
295 State of theory * (ctyp Inttab.table * int) * (cterm Inttab.table * int) * (thm Inttab.table * int) |
296 State of theory * (ctyp Inttab.table * int) * (cterm Inttab.table * int) * (thm Inttab.table * int) |
296 |
297 |
303 (case Inttab.lookup tab (Int.abs i) of |
304 (case Inttab.lookup tab (Int.abs i) of |
304 NONE => raise Fail "get: lookup failed" |
305 NONE => raise Fail "get: lookup failed" |
305 | SOME res => (res, (if i < 0 then Inttab.delete (Int.abs i) tab else tab, no)))) |
306 | SOME res => (res, (if i < 0 then Inttab.delete (Int.abs i) tab else tab, no)))) |
306 |
307 |
307 fun get_theory (State (thy, _, _, _)) = thy; |
308 fun get_theory (State (thy, _, _, _)) = thy; |
308 fun map_theory f (State (thy, a, b, c)) = State (f thy, a, b, c); |
309 val theory = `get_theory; |
309 fun map_theory_result f (State (thy, a, b, c)) = |
310 fun theory_op f (State (thy, a, b, c)) = |
310 let val (res, thy') = f thy in (res, State (thy', a, b, c)) end; |
311 let val (res, thy') = f thy in (res, State (thy', a, b, c)) end; |
311 |
|
312 fun ctyp_of (State (thy, _, _, _)) = Thm.global_ctyp_of thy; |
|
313 fun cterm_of (State (thy, _, _, _)) = Thm.global_cterm_of thy; |
|
314 |
312 |
315 fun typ i (State (thy, a, b, c)) = let val (i, a') = get a i in (i, State (thy, a', b, c)) end |
313 fun typ i (State (thy, a, b, c)) = let val (i, a') = get a i in (i, State (thy, a', b, c)) end |
316 fun term i (State (thy, a, b, c)) = let val (i, b') = get b i in (i, State (thy, a, b', c)) end |
314 fun term i (State (thy, a, b, c)) = let val (i, b') = get b i in (i, State (thy, a, b', c)) end |
317 fun thm i (State (thy, a, b, c)) = let val (i, c') = get c i in (i, State (thy, a, b, c')) end |
315 fun thm i (State (thy, a, b, c)) = let val (i, c') = get c i in (i, State (thy, a, b, c')) end |
318 |
316 |
319 fun set (tab, no) v = (Inttab.update_new (no + 1, v) tab, no + 1) |
317 fun set (tab, no) v = (Inttab.update_new (no + 1, v) tab, no + 1) |
320 fun set_typ ty (State (thy, a, b, c)) = State (thy, set a ty, b, c) |
318 fun set_typ ty (State (thy, a, b, c)) = State (thy, set a ty, b, c) |
321 fun set_term tm (State (thy, a, b, c)) = State (thy, a, set b tm, c) |
319 fun set_term tm (State (thy, a, b, c)) = State (thy, a, set b tm, c) |
322 fun set_thm th (State (thy, a, b, c)) = State (thy, a, b, set c th) |
320 fun set_thm th (State (thy, a, b, c)) = State (thy, a, b, set c th) |
323 |
321 |
324 fun last_thm (State (_, _, _, (tab, no))) = |
322 fun get_thm name (state as State (thy, _, _, _)) = |
325 case Inttab.lookup tab no of |
323 (freeze thy (Global_Theory.get_thm thy name), state); |
326 NONE => raise Fail "last_thm: lookup failed" |
324 |
327 | SOME th => th |
325 fun store_last_thm binding (State (thy, a, b, c as (tab, no))) = |
328 |
326 let |
329 fun list_last (x :: y :: zs) = apfst (fn t => x :: y :: t) (list_last zs) |
327 val th0 = |
330 | list_last [x] = ([], x) |
328 (case Inttab.lookup tab no of |
331 | list_last [] = raise Fail "list_last: empty" |
329 NONE => raise Fail ("No thm " ^ string_of_int no) |
332 |
330 | SOME th => th) |
333 fun pair_list (x :: y :: zs) = ((x, y) :: pair_list zs) |
|
334 | pair_list [] = [] |
|
335 | pair_list _ = raise Fail "pair_list: odd list length" |
|
336 |
|
337 fun store_thm binding th0 thy = |
|
338 let |
|
339 val ctxt = Proof_Context.init_global thy |
331 val ctxt = Proof_Context.init_global thy |
340 val th = Drule.export_without_context_open th0 |
332 val th = Drule.export_without_context_open th0 |
341 val tvs = Term.add_tvars (Thm.prop_of th) [] |
333 val tvs = Term.add_tvars (Thm.prop_of th) [] |
342 val tns = map (fn (_, _) => "'") tvs |
334 val tns = map (fn (_, _) => "'") tvs |
343 val nms = Name.variants (Variable.names_of ctxt) tns |
335 val nms = Name.variants (Variable.names_of ctxt) tns |
344 val vs = map TVar ((nms ~~ (map (snd o fst) tvs)) ~~ (map snd tvs)) |
336 val vs = map TVar ((nms ~~ (map (snd o fst) tvs)) ~~ (map snd tvs)) |
345 val th' = Thm.instantiate (TVars.make (tvs ~~ map (Thm.ctyp_of ctxt) vs), Vars.empty) th |
337 val th' = Thm.instantiate (TVars.make (tvs ~~ map (Thm.ctyp_of ctxt) vs), Vars.empty) th |
346 in |
338 val thy' = #2 (Global_Theory.add_thm ((binding, th'), []) thy) |
347 snd (Global_Theory.add_thm ((binding, th'), []) thy) |
339 in State (thy', a, b, c) end |
348 end |
340 |
|
341 fun list_last (x :: y :: zs) = apfst (fn t => x :: y :: t) (list_last zs) |
|
342 | list_last [x] = ([], x) |
|
343 | list_last [] = raise Fail "list_last: empty" |
|
344 |
|
345 fun pair_list (x :: y :: zs) = ((x, y) :: pair_list zs) |
|
346 | pair_list [] = [] |
|
347 | pair_list _ = raise Fail "pair_list: odd list length" |
349 |
348 |
350 fun parse_line s = |
349 fun parse_line s = |
351 (case String.tokens (fn x => x = #"\n" orelse x = #" ") s of |
350 (case String.tokens (fn x => x = #"\n" orelse x = #" ") s of |
352 [] => raise Fail "parse_line: empty" |
351 [] => raise Fail "parse_line: empty" |
353 | cmd :: args => |
352 | cmd :: args => |
366 term t #>> Thm.apply \<^cterm>\<open>Trueprop\<close> #>> Skip_Proof.make_thm_cterm #-> set_thm |
365 term t #>> Thm.apply \<^cterm>\<open>Trueprop\<close> #>> Skip_Proof.make_thm_cterm #-> set_thm |
367 | process (#"C", [th1, th2]) = thm th1 ##>> thm th2 #>> uncurry comb #-> set_thm |
366 | process (#"C", [th1, th2]) = thm th1 ##>> thm th2 #>> uncurry comb #-> set_thm |
368 | process (#"T", [th1, th2]) = thm th1 ##>> thm th2 #>> uncurry trans #-> set_thm |
367 | process (#"T", [th1, th2]) = thm th1 ##>> thm th2 #>> uncurry trans #-> set_thm |
369 | process (#"E", [th1, th2]) = thm th1 ##>> thm th2 #>> uncurry eq_mp #-> set_thm |
368 | process (#"E", [th1, th2]) = thm th1 ##>> thm th2 #>> uncurry eq_mp #-> set_thm |
370 | process (#"D", [th1, th2]) = thm th1 ##>> thm th2 #>> uncurry deduct #-> set_thm |
369 | process (#"D", [th1, th2]) = thm th1 ##>> thm th2 #>> uncurry deduct #-> set_thm |
371 | process (#"L", [t, th]) = term t ##>> (fn ti => thm th ti) #>> uncurry abs #-> set_thm |
370 | process (#"L", [t, th]) = term t ##>> thm th #>> uncurry abs #-> set_thm |
372 | process (#"M", [s]) = (fn state => |
371 | process (#"M", [s]) = get_thm s #-> set_thm |
373 let |
372 | process (#"Q", args) = |
374 val thy = get_theory state |
373 list_last args |> (fn (tys, th) => |
375 val th = Global_Theory.get_thm thy s |
374 thm th #-> (fn th => fold_map typ tys #-> (fn tys => |
376 in |
375 set_thm (inst_type (pair_list tys) th)))) |
377 set_thm (freeze thy th) state |
376 | process (#"S", args) = |
378 end) |
377 list_last args |> (fn (ts, th) => |
379 | process (#"Q", l) = (fn state => |
378 thm th #-> (fn th => fold_map term ts #-> (fn ts => |
380 let |
379 set_thm (inst (pair_list ts) th)))) |
381 val (tys, th) = list_last l |
380 | process (#"F", [name, t]) = |
382 val (th, state1) = thm th state |
381 term t #-> (fn t => theory_op (def (make_name name) t) #-> set_thm) |
383 val (tys, state2) = fold_map typ tys state1 |
382 | process (#"F", [name]) = theory #-> (fn thy => set_thm (mdef thy name)) |
384 in |
383 | process (#"Y", [name, abs, rep, t1, t2, th]) = |
385 set_thm (inst_type (pair_list tys) th) state2 |
384 thm th #-> (fn th => term t1 #-> (fn t1 => term t2 #-> (fn t2 => |
386 end) |
385 theory_op (tydef name abs rep t1 t2 th) #-> set_thm))) |
387 | process (#"S", l) = (fn state => |
386 | process (#"Y", [name, _, _]) = theory #-> (fn thy => set_thm (mtydef thy name)) |
388 let |
387 | process (#"t", [n]) = theory #-> (fn thy => set_typ (make_tfree thy n)) |
389 val (tms, th) = list_last l |
388 | process (#"a", n :: tys) = theory #-> (fn thy => |
390 val (th, state1) = thm th state |
389 fold_map typ tys #-> (fn tys => set_typ (make_type thy (n, tys)))) |
391 val (tms, state2) = fold_map term tms state1 |
390 | process (#"v", [n, ty]) = typ ty #>> curry make_free n #-> set_term |
392 in |
391 | process (#"c", [n, ty]) = theory #-> (fn thy => |
393 set_thm (inst (pair_list tms) th) state2 |
392 typ ty #>> curry (make_const thy) n #-> set_term) |
394 end) |
|
395 | process (#"F", [name, t]) = (fn state => |
|
396 let |
|
397 val (tm, state1) = term t state |
|
398 in |
|
399 state1 |
|
400 |> map_theory_result (def (make_name name) tm) |
|
401 |-> set_thm |
|
402 end) |
|
403 | process (#"F", [name]) = (fn state => set_thm (mdef (get_theory state) name) state) |
|
404 | process (#"Y", [name, absname, repname, t1, t2, th]) = (fn state => |
|
405 let |
|
406 val (th, state1) = thm th state |
|
407 val (t1, state2) = term t1 state1 |
|
408 val (t2, state3) = term t2 state2 |
|
409 in |
|
410 state3 |
|
411 |> map_theory_result (tydef name absname repname t1 t2 th) |
|
412 |-> set_thm |
|
413 end) |
|
414 | process (#"Y", [name, _, _]) = (fn state => set_thm (mtydef (get_theory state) name) state) |
|
415 | process (#"t", [n]) = (fn state => set_typ (ctyp_of state (make_tfree n)) state) |
|
416 | process (#"a", n :: l) = (fn state => |
|
417 fold_map typ l state |
|
418 |>> (fn tys => ctyp_of state (make_type (get_theory state) (n, map Thm.typ_of tys))) |
|
419 |-> set_typ) |
|
420 | process (#"v", [n, ty]) = (fn state => |
|
421 typ ty state |>> (fn ty => cterm_of state (make_free (n, Thm.typ_of ty))) |-> set_term) |
|
422 | process (#"c", [n, ty]) = (fn state => |
|
423 typ ty state |>> (fn ty => |
|
424 cterm_of state (make_const (get_theory state) (n, Thm.typ_of ty))) |-> set_term) |
|
425 | process (#"f", [t1, t2]) = term t1 ##>> term t2 #>> uncurry Thm.apply #-> set_term |
393 | process (#"f", [t1, t2]) = term t1 ##>> term t2 #>> uncurry Thm.apply #-> set_term |
426 | process (#"l", [t1, t2]) = term t1 ##>> term t2 #>> uncurry Thm.lambda #-> set_term |
394 | process (#"l", [t1, t2]) = term t1 ##>> term t2 #>> uncurry Thm.lambda #-> set_term |
427 | process (#"+", [s]) = (fn state => |
395 | process (#"+", [s]) = store_last_thm (Binding.name (make_name s)) |
428 map_theory (store_thm (Binding.name (make_name s)) (last_thm state)) state) |
|
429 | process (c, _) = raise Fail ("process: unknown command: " ^ String.str c) |
396 | process (c, _) = raise Fail ("process: unknown command: " ^ String.str c) |
430 in |
397 in |
431 process (parse_line str) |
398 process (parse_line str) |
432 end |
399 end |
433 |
400 |