src/HOL/HOLCF/Tools/cont_proc.ML
author wenzelm
Mon, 11 Sep 2023 19:30:48 +0200
changeset 78659 b5f3d1051b13
parent 74305 28a582aa25dd
child 78793 2cb027b95188
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 42083
diff changeset
     1
(*  Title:      HOL/HOLCF/Tools/cont_proc.ML
23152
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
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 51717
diff changeset
    10
  val cont_tac: Proof.context -> int -> tactic
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    11
  val cont_proc: simproc
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    12
  val setup: theory -> theory
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    13
end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    14
41296
6aaf80ea9715 switch to transparent ascription, to avoid warning messages
huffman
parents: 40832
diff changeset
    15
structure ContProc : CONT_PROC =
23152
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
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    18
(** theory context references **)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    19
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    20
val cont_K = @{thm cont_const}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    21
val cont_I = @{thm cont_id}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    22
val cont_A = @{thm cont2cont_APP}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    23
val cont_L = @{thm cont2cont_LAM}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    24
val cont_R = @{thm cont_Rep_cfun2}
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    25
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    26
(* checks whether a term is written entirely in the LCF sublanguage *)
74305
28a582aa25dd more antiquotations;
wenzelm
parents: 70494
diff changeset
    27
fun is_lcf_term \<^Const_>\<open>Rep_cfun _ _ for t u\<close> = is_lcf_term t andalso is_lcf_term u
28a582aa25dd more antiquotations;
wenzelm
parents: 70494
diff changeset
    28
  | is_lcf_term \<^Const_>\<open>Abs_cfun _ _ for \<open>Abs (_, _, t)\<close>\<close> = is_lcf_term t
28a582aa25dd more antiquotations;
wenzelm
parents: 70494
diff changeset
    29
  | is_lcf_term \<^Const_>\<open>Abs_cfun _ _ for t\<close> = is_lcf_term (Term.incr_boundvars 1 t $ Bound 0)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    30
  | is_lcf_term (Bound _) = true
42083
e1209fc7ecdc added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
wenzelm
parents: 41296
diff changeset
    31
  | is_lcf_term t = not (Term.is_open t)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    32
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    33
(*
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    34
  efficiently generates a cont thm for every LAM abstraction in a term,
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    35
  using forward proof and reusing common subgoals
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    36
*)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    37
local
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    38
  fun var 0 = [SOME cont_I]
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    39
    | var n = NONE :: var (n-1)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    40
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    41
  fun k NONE     = cont_K
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    42
    | k (SOME x) = x
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    43
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    44
  fun ap NONE NONE = NONE
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    45
    | ap x    y    = SOME (k y RS (k x RS cont_A))
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    46
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    47
  fun zip []      []      = []
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    48
    | zip []      (y::ys) = (ap NONE y   ) :: zip [] ys
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    49
    | zip (x::xs) []      = (ap x    NONE) :: zip xs []
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    50
    | zip (x::xs) (y::ys) = (ap x    y   ) :: zip xs ys
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    51
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    52
  fun lam [] = ([], cont_K)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    53
    | lam (x::ys) =
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    54
    let
29371
bab4e907d881 use Thm.close_derivation instead of standard
huffman
parents: 29048
diff changeset
    55
      (* should use "close_derivation" for thms that are used multiple times *)
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    56
      (* it seems to allow for sharing in explicit proof objects *)
70494
41108e3e9ca5 formal position for PThm nodes;
wenzelm
parents: 69597
diff changeset
    57
      val x' = Thm.close_derivation \<^here> (k x)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    58
      val Lx = x' RS cont_L
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    59
    in (map (fn y => SOME (k y RS Lx)) ys, x') end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    60
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    61
  (* first list: cont thm for each dangling bound variable *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    62
  (* second list: cont thm for each LAM in t *)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    63
  (* if b = false, only return cont thm for outermost LAMs *)
74305
28a582aa25dd more antiquotations;
wenzelm
parents: 70494
diff changeset
    64
  fun cont_thms1 b \<^Const_>\<open>Rep_cfun _ _ for f t\<close> =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    65
    let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    66
      val (cs1,ls1) = cont_thms1 b f
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    67
      val (cs2,ls2) = cont_thms1 b t
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    68
    in (zip cs1 cs2, if b then ls1 @ ls2 else []) end
74305
28a582aa25dd more antiquotations;
wenzelm
parents: 70494
diff changeset
    69
    | cont_thms1 b \<^Const_>\<open>Abs_cfun _ _ for \<open>Abs (_, _, t)\<close>\<close> =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    70
    let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    71
      val (cs, ls) = cont_thms1 b t
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    72
      val (cs', l) = lam cs
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    73
    in (cs', l::ls) end
74305
28a582aa25dd more antiquotations;
wenzelm
parents: 70494
diff changeset
    74
    | cont_thms1 b \<^Const_>\<open>Abs_cfun _ _ for t\<close> =
29373
6a19d9f6021d make cont_proc handle eta-contracted terms
huffman
parents: 29372
diff changeset
    75
    let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    76
      val t' = Term.incr_boundvars 1 t $ Bound 0
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    77
      val (cs, ls) = cont_thms1 b t'
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    78
      val (cs', l) = lam cs
29373
6a19d9f6021d make cont_proc handle eta-contracted terms
huffman
parents: 29372
diff changeset
    79
    in (cs', l::ls) end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    80
    | cont_thms1 _ (Bound n) = (var n, [])
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    81
    | cont_thms1 _ _ = ([], [])
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    82
in
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    83
  (* precondition: is_lcf_term t = true *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    84
  fun cont_thms t = snd (cont_thms1 false t)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    85
  fun all_cont_thms t = snd (cont_thms1 true t)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    86
end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    87
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    88
(*
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    89
  Given the term "cont f", the procedure tries to construct the
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    90
  theorem "cont f == True". If this theorem cannot be completely
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    91
  solved by the introduction rules, then the procedure returns a
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    92
  conditional rewrite rule with the unsolved subgoals as premises.
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    93
*)
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    94
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 51717
diff changeset
    95
fun cont_tac ctxt =
29048
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
    96
  let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    97
    val rules = [cont_K, cont_I, cont_R, cont_A, cont_L]
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
    98
  
29662
c8c67557f187 removed use of prev_cont_thms reference
huffman
parents: 29373
diff changeset
    99
    fun new_cont_tac f' i =
29048
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   100
      case all_cont_thms f' of
29662
c8c67557f187 removed use of prev_cont_thms reference
huffman
parents: 29373
diff changeset
   101
        [] => no_tac
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59643
diff changeset
   102
      | (c::_) => resolve_tac ctxt [c] i
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   103
74305
28a582aa25dd more antiquotations;
wenzelm
parents: 70494
diff changeset
   104
    fun cont_tac_of_term \<^Const_>\<open>cont _ _ for f\<close> =
29048
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   105
      let
74305
28a582aa25dd more antiquotations;
wenzelm
parents: 70494
diff changeset
   106
        val f' = \<^Const>\<open>Abs_cfun dummyT dummyT for f\<close>
29048
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   107
      in
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   108
        if is_lcf_term f'
29662
c8c67557f187 removed use of prev_cont_thms reference
huffman
parents: 29373
diff changeset
   109
        then new_cont_tac f'
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 51717
diff changeset
   110
        else REPEAT_ALL_NEW (resolve_tac ctxt rules)
29048
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   111
      end
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   112
      | cont_tac_of_term _ = K no_tac
29048
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   113
  in
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   114
    SUBGOAL (fn (t, i) =>
0735d0f89172 implement cont_proc theorem cache using theory data
huffman
parents: 29047
diff changeset
   115
      cont_tac_of_term (HOLogic.dest_Trueprop t) i)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   116
  end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   117
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   118
local
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   119
  fun solve_cont ctxt ct =
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   120
    let
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   121
      val t = Thm.term_of ct
60801
7664e0916eec tuned signature;
wenzelm
parents: 60754
diff changeset
   122
      val tr = Thm.instantiate' [] [SOME (Thm.cterm_of ctxt t)] @{thm Eq_TrueI}
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 51717
diff changeset
   123
    in Option.map fst (Seq.pull (cont_tac ctxt 1 tr)) end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   124
in
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   125
  val cont_proc =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62913
diff changeset
   126
    Simplifier.make_simproc \<^context> "cont_proc"
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62913
diff changeset
   127
     {lhss = [\<^term>\<open>cont f\<close>], proc = K solve_cont}
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   128
end
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   129
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   130
val setup = map_theory_simpset (fn ctxt => ctxt addsimprocs [cont_proc])
23152
9497234a2743 moved HOLCF tools to canonical place;
wenzelm
parents:
diff changeset
   131
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   132
end