src/HOLCF/Tools/cont_proc.ML
author huffman
Tue, 06 Jan 2009 09:02:18 -0800
changeset 29372 df457e0d9a55
parent 29371 bab4e907d881
child 29373 6a19d9f6021d
permissions -rw-r--r--
implement is_closed_term using Term.loose_bvar
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOLCF/Tools/cont_proc.ML
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     2
    Author:     Brian Huffman
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     3
*)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     4
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     5
signature CONT_PROC =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     6
sig
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     7
  val is_lcf_term: term -> bool
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     8
  val cont_thms: term -> thm list
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
     9
  val all_cont_thms: term -> thm list
29048
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
    10
  val cont_tac: thm list ref -> int -> tactic
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    11
  val cont_proc: theory -> simproc
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    12
  val setup: theory -> theory
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    13
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    14
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    15
structure ContProc: CONT_PROC =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    16
struct
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    17
29048
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
    18
structure ContProcData = TheoryDataFun
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
    19
(
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
    20
  type T = thm list ref;
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
    21
  val empty = ref [];
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
    22
  val copy = I;
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
    23
  val extend = I;
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
    24
  fun merge _ _ = ref [];
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
    25
)
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
    26
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    27
(** theory context references **)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    28
29047
059bdb9813c6 use ML antiquotations
huffman
parents: 26496
diff changeset
    29
val cont_K = @{thm cont_const};
059bdb9813c6 use ML antiquotations
huffman
parents: 26496
diff changeset
    30
val cont_I = @{thm cont_id};
059bdb9813c6 use ML antiquotations
huffman
parents: 26496
diff changeset
    31
val cont_A = @{thm cont2cont_Rep_CFun};
059bdb9813c6 use ML antiquotations
huffman
parents: 26496
diff changeset
    32
val cont_L = @{thm cont2cont_LAM};
059bdb9813c6 use ML antiquotations
huffman
parents: 26496
diff changeset
    33
val cont_R = @{thm cont_Rep_CFun2};
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    34
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    35
(* checks whether a term contains no dangling bound variables *)
29372
df457e0d9a55 implement is_closed_term using Term.loose_bvar
huffman
parents: 29371
diff changeset
    36
fun is_closed_term t = not (Term.loose_bvar (t, 0));
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    37
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    38
(* checks whether a term is written entirely in the LCF sublanguage *)
29047
059bdb9813c6 use ML antiquotations
huffman
parents: 26496
diff changeset
    39
fun is_lcf_term (Const (@{const_name Rep_CFun}, _) $ t $ u) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    40
      is_lcf_term t andalso is_lcf_term u
29047
059bdb9813c6 use ML antiquotations
huffman
parents: 26496
diff changeset
    41
  | is_lcf_term (Const (@{const_name Abs_CFun}, _) $ Abs (_, _, t)) =
059bdb9813c6 use ML antiquotations
huffman
parents: 26496
diff changeset
    42
      is_lcf_term t
059bdb9813c6 use ML antiquotations
huffman
parents: 26496
diff changeset
    43
  | is_lcf_term (Const (@{const_name Abs_CFun}, _) $ _) = false
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    44
  | is_lcf_term (Bound _) = true
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    45
  | is_lcf_term t = is_closed_term t;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    46
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    47
(*
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    48
  efficiently generates a cont thm for every LAM abstraction in a term,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    49
  using forward proof and reusing common subgoals
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    50
*)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    51
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    52
  fun var 0 = [SOME cont_I]
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    53
    | var n = NONE :: var (n-1);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    54
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    55
  fun k NONE     = cont_K
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    56
    | k (SOME x) = x;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    57
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    58
  fun ap NONE NONE = NONE
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    59
    | ap x    y    = SOME (k y RS (k x RS cont_A));
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    60
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    61
  fun zip []      []      = []
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    62
    | zip []      (y::ys) = (ap NONE y   ) :: zip [] ys
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    63
    | zip (x::xs) []      = (ap x    NONE) :: zip xs []
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    64
    | zip (x::xs) (y::ys) = (ap x    y   ) :: zip xs ys
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    65
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    66
  fun lam [] = ([], cont_K)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    67
    | lam (x::ys) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    68
    let
29371
bab4e907d881 use Thm.close_derivation instead of standard
huffman
parents: 29048
diff changeset
    69
      (* should use "close_derivation" for thms that are used multiple times *)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    70
      (* it seems to allow for sharing in explicit proof objects *)
29371
bab4e907d881 use Thm.close_derivation instead of standard
huffman
parents: 29048
diff changeset
    71
      val x' = Thm.close_derivation (k x);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    72
      val Lx = x' RS cont_L;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    73
    in (map (fn y => SOME (k y RS Lx)) ys, x') end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    74
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    75
  (* first list: cont thm for each dangling bound variable *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    76
  (* second list: cont thm for each LAM in t *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    77
  (* if b = false, only return cont thm for outermost LAMs *)
29047
059bdb9813c6 use ML antiquotations
huffman
parents: 26496
diff changeset
    78
  fun cont_thms1 b (Const (@{const_name Rep_CFun}, _) $ f $ t) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    79
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    80
      val (cs1,ls1) = cont_thms1 b f;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    81
      val (cs2,ls2) = cont_thms1 b t;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    82
    in (zip cs1 cs2, if b then ls1 @ ls2 else []) end
29047
059bdb9813c6 use ML antiquotations
huffman
parents: 26496
diff changeset
    83
    | cont_thms1 b (Const (@{const_name Abs_CFun}, _) $ Abs (_, _, t)) =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    84
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    85
      val (cs, ls) = cont_thms1 b t;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    86
      val (cs', l) = lam cs;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    87
    in (cs', l::ls) end
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    88
    | cont_thms1 _ (Bound n) = (var n, [])
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    89
    | cont_thms1 _ _ = ([], []);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    90
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    91
  (* precondition: is_lcf_term t = true *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    92
  fun cont_thms t = snd (cont_thms1 false t);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    93
  fun all_cont_thms t = snd (cont_thms1 true t);
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    94
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    95
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    96
(*
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    97
  Given the term "cont f", the procedure tries to construct the
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    98
  theorem "cont f == True". If this theorem cannot be completely
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    99
  solved by the introduction rules, then the procedure returns a
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   100
  conditional rewrite rule with the unsolved subgoals as premises.
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   101
*)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   102
29048
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   103
fun cont_tac prev_cont_thms =
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   104
  let
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   105
    val rules = [cont_K, cont_I, cont_R, cont_A, cont_L];
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   106
  
29048
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   107
    fun old_cont_tac i thm =
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   108
      case !prev_cont_thms of
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   109
        [] => no_tac thm
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   110
      | (c::cs) => (prev_cont_thms := cs; rtac c i thm);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   111
29048
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   112
    fun new_cont_tac f' i thm =
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   113
      case all_cont_thms f' of
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   114
        [] => no_tac thm
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   115
      | (c::cs) => (prev_cont_thms := cs; rtac c i thm);
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   116
29048
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   117
    fun cont_tac_of_term (Const (@{const_name cont}, _) $ f) =
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   118
      let
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   119
        val f' = Const (@{const_name Abs_CFun}, dummyT) $ f;
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   120
      in
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   121
        if is_lcf_term f'
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   122
        then old_cont_tac ORELSE' new_cont_tac f'
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   123
        else REPEAT_ALL_NEW (resolve_tac rules)
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   124
      end
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   125
      | cont_tac_of_term _ = K no_tac;
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   126
  in
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   127
    SUBGOAL (fn (t, i) =>
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   128
      cont_tac_of_term (HOLogic.dest_Trueprop t) i)
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   129
  end;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   130
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   131
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   132
  fun solve_cont thy _ t =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   133
    let
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   134
      val tr = instantiate' [] [SOME (cterm_of thy t)] Eq_TrueI;
29048
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   135
      val prev_thms = ContProcData.get thy
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   136
    in Option.map fst (Seq.pull (cont_tac prev_thms 1 tr)) end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   137
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   138
  fun cont_proc thy =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   139
    Simplifier.simproc thy "cont_proc" ["cont f"] solve_cont;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   140
end;
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   141
26496
49ae9456eba9 purely functional setup of claset/simpset/clasimpset;
wenzelm
parents: 24043
diff changeset
   142
fun setup thy = Simplifier.map_simpset (fn ss => ss addsimprocs [cont_proc thy]) thy;
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   143
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   144
end;