| author | wenzelm | 
| Wed, 27 Mar 2024 13:23:15 +0100 | |
| changeset 80026 | a03a7d4b82f8 | 
| parent 78800 | 0b3700d31758 | 
| permissions | -rw-r--r-- | 
| 42151 | 1 | (* Title: HOL/HOLCF/Tools/cont_proc.ML | 
| 23152 | 2 | Author: Brian Huffman | 
| 3 | *) | |
| 4 | ||
| 5 | signature CONT_PROC = | |
| 6 | sig | |
| 7 | val is_lcf_term: term -> bool | |
| 8 | val cont_thms: term -> thm list | |
| 9 | val all_cont_thms: term -> thm list | |
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
51717diff
changeset | 10 | val cont_tac: Proof.context -> int -> tactic | 
| 78800 | 11 | val cont_proc: Simplifier.proc | 
| 40832 | 12 | end | 
| 23152 | 13 | |
| 41296 
6aaf80ea9715
switch to transparent ascription, to avoid warning messages
 huffman parents: 
40832diff
changeset | 14 | structure ContProc : CONT_PROC = | 
| 23152 | 15 | struct | 
| 16 | ||
| 17 | (** theory context references **) | |
| 18 | ||
| 40832 | 19 | val cont_K = @{thm cont_const}
 | 
| 20 | val cont_I = @{thm cont_id}
 | |
| 21 | val cont_A = @{thm cont2cont_APP}
 | |
| 22 | val cont_L = @{thm cont2cont_LAM}
 | |
| 23 | val cont_R = @{thm cont_Rep_cfun2}
 | |
| 23152 | 24 | |
| 25 | (* checks whether a term is written entirely in the LCF sublanguage *) | |
| 74305 | 26 | fun is_lcf_term \<^Const_>\<open>Rep_cfun _ _ for t u\<close> = is_lcf_term t andalso is_lcf_term u | 
| 27 | | is_lcf_term \<^Const_>\<open>Abs_cfun _ _ for \<open>Abs (_, _, t)\<close>\<close> = is_lcf_term t | |
| 28 | | is_lcf_term \<^Const_>\<open>Abs_cfun _ _ for t\<close> = is_lcf_term (Term.incr_boundvars 1 t $ Bound 0) | |
| 23152 | 29 | | 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: 
41296diff
changeset | 30 | | is_lcf_term t = not (Term.is_open t) | 
| 23152 | 31 | |
| 32 | (* | |
| 33 | efficiently generates a cont thm for every LAM abstraction in a term, | |
| 34 | using forward proof and reusing common subgoals | |
| 35 | *) | |
| 36 | local | |
| 37 | fun var 0 = [SOME cont_I] | |
| 40832 | 38 | | var n = NONE :: var (n-1) | 
| 23152 | 39 | |
| 40 | fun k NONE = cont_K | |
| 40832 | 41 | | k (SOME x) = x | 
| 23152 | 42 | |
| 43 | fun ap NONE NONE = NONE | |
| 40832 | 44 | | ap x y = SOME (k y RS (k x RS cont_A)) | 
| 23152 | 45 | |
| 46 | fun zip [] [] = [] | |
| 47 | | zip [] (y::ys) = (ap NONE y ) :: zip [] ys | |
| 48 | | zip (x::xs) [] = (ap x NONE) :: zip xs [] | |
| 49 | | zip (x::xs) (y::ys) = (ap x y ) :: zip xs ys | |
| 50 | ||
| 51 | fun lam [] = ([], cont_K) | |
| 52 | | lam (x::ys) = | |
| 53 | let | |
| 29371 | 54 | (* should use "close_derivation" for thms that are used multiple times *) | 
| 23152 | 55 | (* it seems to allow for sharing in explicit proof objects *) | 
| 70494 | 56 | val x' = Thm.close_derivation \<^here> (k x) | 
| 40832 | 57 | val Lx = x' RS cont_L | 
| 58 | in (map (fn y => SOME (k y RS Lx)) ys, x') end | |
| 23152 | 59 | |
| 60 | (* first list: cont thm for each dangling bound variable *) | |
| 61 | (* second list: cont thm for each LAM in t *) | |
| 62 | (* if b = false, only return cont thm for outermost LAMs *) | |
| 74305 | 63 | fun cont_thms1 b \<^Const_>\<open>Rep_cfun _ _ for f t\<close> = | 
| 23152 | 64 | let | 
| 40832 | 65 | val (cs1,ls1) = cont_thms1 b f | 
| 66 | val (cs2,ls2) = cont_thms1 b t | |
| 23152 | 67 | in (zip cs1 cs2, if b then ls1 @ ls2 else []) end | 
| 74305 | 68 | | cont_thms1 b \<^Const_>\<open>Abs_cfun _ _ for \<open>Abs (_, _, t)\<close>\<close> = | 
| 23152 | 69 | let | 
| 40832 | 70 | val (cs, ls) = cont_thms1 b t | 
| 71 | val (cs', l) = lam cs | |
| 23152 | 72 | in (cs', l::ls) end | 
| 74305 | 73 | | cont_thms1 b \<^Const_>\<open>Abs_cfun _ _ for t\<close> = | 
| 29373 | 74 | let | 
| 40832 | 75 | val t' = Term.incr_boundvars 1 t $ Bound 0 | 
| 76 | val (cs, ls) = cont_thms1 b t' | |
| 77 | val (cs', l) = lam cs | |
| 29373 | 78 | in (cs', l::ls) end | 
| 23152 | 79 | | cont_thms1 _ (Bound n) = (var n, []) | 
| 40832 | 80 | | cont_thms1 _ _ = ([], []) | 
| 23152 | 81 | in | 
| 82 | (* precondition: is_lcf_term t = true *) | |
| 40832 | 83 | fun cont_thms t = snd (cont_thms1 false t) | 
| 84 | fun all_cont_thms t = snd (cont_thms1 true t) | |
| 85 | end | |
| 23152 | 86 | |
| 87 | (* | |
| 88 | Given the term "cont f", the procedure tries to construct the | |
| 89 | theorem "cont f == True". If this theorem cannot be completely | |
| 90 | solved by the introduction rules, then the procedure returns a | |
| 91 | conditional rewrite rule with the unsolved subgoals as premises. | |
| 92 | *) | |
| 93 | ||
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
51717diff
changeset | 94 | fun cont_tac ctxt = | 
| 29048 
0735d0f89172
implement cont_proc theorem cache using theory data
 huffman parents: 
29047diff
changeset | 95 | let | 
| 40832 | 96 | val rules = [cont_K, cont_I, cont_R, cont_A, cont_L] | 
| 23152 | 97 | |
| 29662 | 98 | fun new_cont_tac f' i = | 
| 29048 
0735d0f89172
implement cont_proc theorem cache using theory data
 huffman parents: 
29047diff
changeset | 99 | case all_cont_thms f' of | 
| 29662 | 100 | [] => no_tac | 
| 60754 | 101 | | (c::_) => resolve_tac ctxt [c] i | 
| 23152 | 102 | |
| 74305 | 103 | fun cont_tac_of_term \<^Const_>\<open>cont _ _ for f\<close> = | 
| 29048 
0735d0f89172
implement cont_proc theorem cache using theory data
 huffman parents: 
29047diff
changeset | 104 | let | 
| 74305 | 105 | val f' = \<^Const>\<open>Abs_cfun dummyT dummyT for f\<close> | 
| 29048 
0735d0f89172
implement cont_proc theorem cache using theory data
 huffman parents: 
29047diff
changeset | 106 | in | 
| 
0735d0f89172
implement cont_proc theorem cache using theory data
 huffman parents: 
29047diff
changeset | 107 | if is_lcf_term f' | 
| 29662 | 108 | then new_cont_tac f' | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
51717diff
changeset | 109 | else REPEAT_ALL_NEW (resolve_tac ctxt rules) | 
| 29048 
0735d0f89172
implement cont_proc theorem cache using theory data
 huffman parents: 
29047diff
changeset | 110 | end | 
| 40832 | 111 | | cont_tac_of_term _ = K no_tac | 
| 29048 
0735d0f89172
implement cont_proc theorem cache using theory data
 huffman parents: 
29047diff
changeset | 112 | in | 
| 
0735d0f89172
implement cont_proc theorem cache using theory data
 huffman parents: 
29047diff
changeset | 113 | SUBGOAL (fn (t, i) => | 
| 
0735d0f89172
implement cont_proc theorem cache using theory data
 huffman parents: 
29047diff
changeset | 114 | cont_tac_of_term (HOLogic.dest_Trueprop t) i) | 
| 40832 | 115 | end | 
| 23152 | 116 | |
| 78793 | 117 | fun cont_proc ctxt ct = | 
| 118 | let | |
| 119 | val t = Thm.term_of ct | |
| 120 |     val tr = Thm.instantiate' [] [SOME (Thm.cterm_of ctxt t)] @{thm Eq_TrueI}
 | |
| 121 | in Option.map fst (Seq.pull (cont_tac ctxt 1 tr)) end | |
| 23152 | 122 | |
| 40832 | 123 | end |