| 23152 |      1 | (*  Title:      HOLCF/Tools/cont_proc.ML
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Brian Huffman
 | 
|  |      4 | *)
 | 
|  |      5 | 
 | 
|  |      6 | signature CONT_PROC =
 | 
|  |      7 | sig
 | 
|  |      8 |   val is_lcf_term: term -> bool
 | 
|  |      9 |   val cont_thms: term -> thm list
 | 
|  |     10 |   val all_cont_thms: term -> thm list
 | 
|  |     11 |   val cont_tac: int -> tactic
 | 
|  |     12 |   val cont_proc: theory -> simproc
 | 
|  |     13 |   val setup: theory -> theory
 | 
|  |     14 | end;
 | 
|  |     15 | 
 | 
|  |     16 | structure ContProc: CONT_PROC =
 | 
|  |     17 | struct
 | 
|  |     18 | 
 | 
|  |     19 | (** theory context references **)
 | 
|  |     20 | 
 | 
|  |     21 | val cont_K = thm "cont_const";
 | 
|  |     22 | val cont_I = thm "cont_id";
 | 
|  |     23 | val cont_A = thm "cont2cont_Rep_CFun";
 | 
|  |     24 | val cont_L = thm "cont2cont_LAM";
 | 
|  |     25 | val cont_R = thm "cont_Rep_CFun2";
 | 
|  |     26 | 
 | 
|  |     27 | (* checks whether a term contains no dangling bound variables *)
 | 
|  |     28 | val is_closed_term =
 | 
|  |     29 |   let
 | 
|  |     30 |     fun bound_less i (t $ u) =
 | 
|  |     31 |           bound_less i t andalso bound_less i u
 | 
|  |     32 |       | bound_less i (Abs (_, _, t)) = bound_less (i+1) t
 | 
|  |     33 |       | bound_less i (Bound n) = n < i
 | 
|  |     34 |       | bound_less i _ = true; (* Const, Free, and Var are OK *)
 | 
|  |     35 |   in bound_less 0 end;
 | 
|  |     36 | 
 | 
|  |     37 | (* checks whether a term is written entirely in the LCF sublanguage *)
 | 
|  |     38 | fun is_lcf_term (Const ("Cfun.Rep_CFun", _) $ t $ u) =
 | 
|  |     39 |       is_lcf_term t andalso is_lcf_term u
 | 
|  |     40 |   | is_lcf_term (Const ("Cfun.Abs_CFun", _) $ Abs (_, _, t)) = is_lcf_term t
 | 
|  |     41 |   | is_lcf_term (Const ("Cfun.Abs_CFun", _) $ _) = false
 | 
|  |     42 |   | is_lcf_term (Bound _) = true
 | 
|  |     43 |   | is_lcf_term t = is_closed_term t;
 | 
|  |     44 | 
 | 
|  |     45 | (*
 | 
|  |     46 |   efficiently generates a cont thm for every LAM abstraction in a term,
 | 
|  |     47 |   using forward proof and reusing common subgoals
 | 
|  |     48 | *)
 | 
|  |     49 | local
 | 
|  |     50 |   fun var 0 = [SOME cont_I]
 | 
|  |     51 |     | var n = NONE :: var (n-1);
 | 
|  |     52 | 
 | 
|  |     53 |   fun k NONE     = cont_K
 | 
|  |     54 |     | k (SOME x) = x;
 | 
|  |     55 | 
 | 
|  |     56 |   fun ap NONE NONE = NONE
 | 
|  |     57 |     | ap x    y    = SOME (k y RS (k x RS cont_A));
 | 
|  |     58 | 
 | 
|  |     59 |   fun zip []      []      = []
 | 
|  |     60 |     | zip []      (y::ys) = (ap NONE y   ) :: zip [] ys
 | 
|  |     61 |     | zip (x::xs) []      = (ap x    NONE) :: zip xs []
 | 
|  |     62 |     | zip (x::xs) (y::ys) = (ap x    y   ) :: zip xs ys
 | 
|  |     63 | 
 | 
|  |     64 |   fun lam [] = ([], cont_K)
 | 
|  |     65 |     | lam (x::ys) =
 | 
|  |     66 |     let
 | 
|  |     67 |       (* should use "standard" for thms that are used multiple times *)
 | 
|  |     68 |       (* it seems to allow for sharing in explicit proof objects *)
 | 
|  |     69 |       val x' = standard (k x);
 | 
|  |     70 |       val Lx = x' RS cont_L;
 | 
|  |     71 |     in (map (fn y => SOME (k y RS Lx)) ys, x') end;
 | 
|  |     72 | 
 | 
|  |     73 |   (* first list: cont thm for each dangling bound variable *)
 | 
|  |     74 |   (* second list: cont thm for each LAM in t *)
 | 
|  |     75 |   (* if b = false, only return cont thm for outermost LAMs *)
 | 
|  |     76 |   fun cont_thms1 b (Const ("Cfun.Rep_CFun", _) $ f $ t) =
 | 
|  |     77 |     let
 | 
|  |     78 |       val (cs1,ls1) = cont_thms1 b f;
 | 
|  |     79 |       val (cs2,ls2) = cont_thms1 b t;
 | 
|  |     80 |     in (zip cs1 cs2, if b then ls1 @ ls2 else []) end
 | 
|  |     81 |     | cont_thms1 b (Const ("Cfun.Abs_CFun", _) $ Abs (_, _, t)) =
 | 
|  |     82 |     let
 | 
|  |     83 |       val (cs, ls) = cont_thms1 b t;
 | 
|  |     84 |       val (cs', l) = lam cs;
 | 
|  |     85 |     in (cs', l::ls) end
 | 
|  |     86 |     | cont_thms1 _ (Bound n) = (var n, [])
 | 
|  |     87 |     | cont_thms1 _ _ = ([], []);
 | 
|  |     88 | in
 | 
|  |     89 |   (* precondition: is_lcf_term t = true *)
 | 
|  |     90 |   fun cont_thms t = snd (cont_thms1 false t);
 | 
|  |     91 |   fun all_cont_thms t = snd (cont_thms1 true t);
 | 
|  |     92 | end;
 | 
|  |     93 | 
 | 
|  |     94 | (*
 | 
|  |     95 |   Given the term "cont f", the procedure tries to construct the
 | 
|  |     96 |   theorem "cont f == True". If this theorem cannot be completely
 | 
|  |     97 |   solved by the introduction rules, then the procedure returns a
 | 
|  |     98 |   conditional rewrite rule with the unsolved subgoals as premises.
 | 
|  |     99 | *)
 | 
|  |    100 | 
 | 
|  |    101 | local
 | 
|  |    102 |   val rules = [cont_K, cont_I, cont_R, cont_A, cont_L];
 | 
|  |    103 |   
 | 
| 24043 |    104 |   (* FIXME proper cache as theory data!? *)
 | 
| 23152 |    105 |   val prev_cont_thms : thm list ref = ref [];
 | 
|  |    106 | 
 | 
| 24043 |    107 |   fun old_cont_tac i thm = CRITICAL (fn () =>
 | 
| 23152 |    108 |     case !prev_cont_thms of
 | 
|  |    109 |       [] => no_tac thm
 | 
| 24043 |    110 |     | (c::cs) => (prev_cont_thms := cs; rtac c i thm));
 | 
| 23152 |    111 | 
 | 
| 24043 |    112 |   fun new_cont_tac f' i thm = CRITICAL (fn () =>
 | 
| 23152 |    113 |     case all_cont_thms f' of
 | 
|  |    114 |       [] => no_tac thm
 | 
| 24043 |    115 |     | (c::cs) => (prev_cont_thms := cs; rtac c i thm));
 | 
| 23152 |    116 | 
 | 
|  |    117 |   fun cont_tac_of_term (Const ("Cont.cont", _) $ f) =
 | 
|  |    118 |     let
 | 
|  |    119 |       val f' = Const ("Cfun.Abs_CFun", dummyT) $ f;
 | 
|  |    120 |     in
 | 
|  |    121 |       if is_lcf_term f'
 | 
|  |    122 |       then old_cont_tac ORELSE' new_cont_tac f'
 | 
|  |    123 |       else REPEAT_ALL_NEW (resolve_tac rules)
 | 
|  |    124 |     end
 | 
|  |    125 |     | cont_tac_of_term _ = K no_tac;
 | 
|  |    126 | in
 | 
|  |    127 |   val cont_tac =
 | 
|  |    128 |     SUBGOAL (fn (t, i) => cont_tac_of_term (HOLogic.dest_Trueprop t) i);
 | 
|  |    129 | end;
 | 
|  |    130 | 
 | 
|  |    131 | local
 | 
|  |    132 |   fun solve_cont thy _ t =
 | 
|  |    133 |     let
 | 
|  |    134 |       val tr = instantiate' [] [SOME (cterm_of thy t)] Eq_TrueI;
 | 
|  |    135 |     in Option.map fst (Seq.pull (cont_tac 1 tr)) end
 | 
|  |    136 | in
 | 
|  |    137 |   fun cont_proc thy =
 | 
|  |    138 |     Simplifier.simproc thy "cont_proc" ["cont f"] solve_cont;
 | 
|  |    139 | end;
 | 
|  |    140 | 
 | 
|  |    141 | val setup =
 | 
|  |    142 |   (fn thy =>
 | 
|  |    143 |     (Simplifier.change_simpset_of thy
 | 
|  |    144 |       (fn ss => ss addsimprocs [cont_proc thy]); thy));
 | 
|  |    145 | 
 | 
|  |    146 | end;
 |