src/HOLCF/cont_proc.ML
author huffman
Wed, 15 Jun 2005 23:35:43 +0200
changeset 16403 294a7864063e
parent 16386 c6f5ade29608
child 16629 91a179d4b0d5
permissions -rw-r--r--
Domain package uses ContProc for beta reduction
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
16386
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
     1
(*  Title:      HOLCF/cont_proc.ML
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
     2
    ID:         $Id$
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
     3
    Author:     Brian Huffman
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
     4
*)
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
     5
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
     6
signature CONT_PROC =
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
     7
sig
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
     8
  val is_lcf_term: term -> bool
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
     9
  val cont_thms: term -> thm list
16403
294a7864063e Domain package uses ContProc for beta reduction
huffman
parents: 16386
diff changeset
    10
  val all_cont_thms: term -> thm list
16386
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    11
  val cont_proc: theory -> simproc
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    12
  val setup: (theory -> theory) list
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    13
end;
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    14
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    15
structure ContProc: CONT_PROC =
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    16
struct
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    17
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    18
(** theory context references **)
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    19
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    20
val cont_K = thm "cont_const";
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    21
val cont_I = thm "cont_id";
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    22
val cont_A = thm "cont2cont_Rep_CFun";
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    23
val cont_L = thm "cont2cont_LAM";
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    24
val cont_R = thm "cont_Rep_CFun2";
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    25
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    26
(* checks whether a term is written entirely in the LCF sublanguage *)
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    27
fun is_lcf_term (Const("Cfun.Rep_CFun",_) $ t $ u) = is_lcf_term t andalso is_lcf_term u
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    28
  | is_lcf_term (Const("Cfun.Abs_CFun",_) $ Abs (_,_,t)) = is_lcf_term t
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    29
  | is_lcf_term (_ $ _) = false
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    30
  | is_lcf_term (Abs _) = false
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    31
  | is_lcf_term _ = true; (* Const, Free, Var, and Bound are OK *)
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    32
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    33
(*
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    34
  efficiently generates a cont thm for every LAM abstraction in a term,
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    35
  using forward proof and reusing common subgoals
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    36
*)
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    37
local
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    38
  fun var 0 = [SOME cont_I]
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    39
    | var n = NONE :: var (n-1);
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    40
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    41
  fun k NONE     = cont_K
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    42
    | k (SOME x) = x;
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    43
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    44
  fun ap NONE NONE = NONE
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    45
    | ap x    y    = SOME (k y RS (k x RS cont_A));
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    46
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    47
  fun zip []      []      = []
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    48
    | zip []      (y::ys) = (ap NONE y   ) :: zip [] ys
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    49
    | zip (x::xs) []      = (ap x    NONE) :: zip xs []
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    50
    | zip (x::xs) (y::ys) = (ap x    y   ) :: zip xs ys
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    51
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    52
  fun lam [] = ([], cont_K)
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    53
    | lam (x::ys) = let
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    54
        (* should use "standard" for thms that are used multiple times *)
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    55
        (* it seems to allow for sharing in explicit proof objects *)
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    56
        val x' = standard (k x);
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    57
        val Lx = x' RS cont_L;
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    58
        in (map (fn y => SOME (k y RS Lx)) ys, x') end;
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    59
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    60
  (* first list: cont thm for each dangling bound variable *)
16403
294a7864063e Domain package uses ContProc for beta reduction
huffman
parents: 16386
diff changeset
    61
  (* second list: cont thm for each LAM in t *)
294a7864063e Domain package uses ContProc for beta reduction
huffman
parents: 16386
diff changeset
    62
  (* if b = false, only return cont thm for outermost LAMs *)
294a7864063e Domain package uses ContProc for beta reduction
huffman
parents: 16386
diff changeset
    63
  fun cont_thms1 b (Const _ $ f $ t) = let
294a7864063e Domain package uses ContProc for beta reduction
huffman
parents: 16386
diff changeset
    64
        val (cs1,ls1) = cont_thms1 b f;
294a7864063e Domain package uses ContProc for beta reduction
huffman
parents: 16386
diff changeset
    65
        val (cs2,ls2) = cont_thms1 b t;
294a7864063e Domain package uses ContProc for beta reduction
huffman
parents: 16386
diff changeset
    66
        in (zip cs1 cs2, if b then ls1 @ ls2 else []) end
294a7864063e Domain package uses ContProc for beta reduction
huffman
parents: 16386
diff changeset
    67
    | cont_thms1 b (Const _ $ Abs (_,_,t)) = let
294a7864063e Domain package uses ContProc for beta reduction
huffman
parents: 16386
diff changeset
    68
        val (cs,ls) = cont_thms1 b t;
16386
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    69
        val (cs',l) = lam cs;
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    70
        in (cs',l::ls) end
16403
294a7864063e Domain package uses ContProc for beta reduction
huffman
parents: 16386
diff changeset
    71
    | cont_thms1 _ (Bound n) = (var n, [])
294a7864063e Domain package uses ContProc for beta reduction
huffman
parents: 16386
diff changeset
    72
    | cont_thms1 _ _ = ([],[]);
16386
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    73
in
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    74
  (* precondition: is_lcf_term t = true *)
16403
294a7864063e Domain package uses ContProc for beta reduction
huffman
parents: 16386
diff changeset
    75
  fun cont_thms t = snd (cont_thms1 false t);
294a7864063e Domain package uses ContProc for beta reduction
huffman
parents: 16386
diff changeset
    76
  fun all_cont_thms t = snd (cont_thms1 true t);
16386
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    77
end;
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    78
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    79
(*
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    80
  Given the term "cont f", the procedure tries to construct the
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    81
  theorem "cont f == True". If this theorem cannot be completely
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    82
  solved by the introduction rules, then the procedure returns a
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    83
  conditional rewrite rule with the unsolved subgoals as premises.
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    84
*)
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    85
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    86
local
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    87
  fun solve_cont sg _ (t as Const("Cont.cont",_) $ f) =
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    88
    let
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    89
      val rules = [cont_K, cont_I, cont_R, cont_A, cont_L];
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    90
      val tr = instantiate' [] [SOME (cterm_of sg t)] Eq_TrueI;
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    91
      val f' = Const("Cfun.Abs_CFun",dummyT) $ f;
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    92
      val tac = if is_lcf_term f'
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    93
        then rtac (hd (cont_thms f')) 1
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    94
        else REPEAT_ALL_NEW (resolve_tac rules) 1;
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    95
    in Option.map fst (Seq.pull (tac tr)) end
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    96
    | solve_cont sg _ _ = NONE;
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    97
in
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    98
  fun cont_proc thy = Simplifier.simproc (Theory.sign_of thy)
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
    99
        "cont_proc" ["cont f"] solve_cont;
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
   100
end;
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
   101
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
   102
val setup =
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
   103
  [fn thy => Simplifier.change_simpset_of (op addsimprocs) [cont_proc thy] thy];
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
   104
c6f5ade29608 moved continuity simproc to a separate file
huffman
parents:
diff changeset
   105
end;