| author | wenzelm | 
| Tue, 07 Jun 2022 19:15:08 +0200 | |
| changeset 75536 | 7cdeed5dc96d | 
| parent 74588 | 3cc363e8bfb2 | 
| child 78807 | f6d2679ab6c1 | 
| permissions | -rw-r--r-- | 
| 28308 | 1 | (* Title: HOL/Statespace/state_fun.ML | 
| 74588 
3cc363e8bfb2
cleanup; add Apple reference
 Norbert Schirmer <nschirmer@apple.com> parents: 
74586diff
changeset | 2 | Author: Norbert Schirmer, TU Muenchen, 2007 | 
| 
3cc363e8bfb2
cleanup; add Apple reference
 Norbert Schirmer <nschirmer@apple.com> parents: 
74586diff
changeset | 3 | Author: Norbert Schirmer, Apple, 2021 | 
| 25171 | 4 | *) | 
| 5 | ||
| 25408 | 6 | signature STATE_FUN = | 
| 7 | sig | |
| 8 | val lookupN : string | |
| 9 | val updateN : string | |
| 25171 | 10 | |
| 29302 | 11 | val mk_constr : theory -> typ -> term | 
| 12 | val mk_destr : theory -> typ -> term | |
| 25408 | 13 | |
| 29302 | 14 | val lookup_simproc : simproc | 
| 15 | val update_simproc : simproc | |
| 16 | val ex_lookup_eq_simproc : simproc | |
| 17 | val ex_lookup_ss : simpset | |
| 18 | val lazy_conj_simproc : simproc | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 19 | val string_eq_simp_tac : Proof.context -> int -> tactic | 
| 74586 
5ac762b53119
generalized component lookup for syntax and distinctness proofs. added some tracing.
 Norbert Schirmer <nschirmer@apple.com> parents: 
74561diff
changeset | 20 | |
| 
5ac762b53119
generalized component lookup for syntax and distinctness proofs. added some tracing.
 Norbert Schirmer <nschirmer@apple.com> parents: 
74561diff
changeset | 21 | val trace_data : Context.generic -> unit | 
| 25408 | 22 | end; | 
| 23 | ||
| 45363 | 24 | structure StateFun: STATE_FUN = | 
| 25171 | 25 | struct | 
| 26 | ||
| 69597 | 27 | val lookupN = \<^const_name>\<open>StateFun.lookup\<close>; | 
| 28 | val updateN = \<^const_name>\<open>StateFun.update\<close>; | |
| 25171 | 29 | |
| 25408 | 30 | val sel_name = HOLogic.dest_string; | 
| 25171 | 31 | |
| 32 | fun mk_name i t = | |
| 33 | (case try sel_name t of | |
| 45363 | 34 | SOME name => name | 
| 35 | | NONE => | |
| 36 | (case t of | |
| 37 | Free (x, _) => x | |
| 38 | | Const (x, _) => x | |
| 39 | | _ => "x" ^ string_of_int i)); | |
| 40 | ||
| 25171 | 41 | local | 
| 32921 | 42 | |
| 39159 | 43 | val conj1_False = @{thm conj1_False};
 | 
| 44 | val conj2_False = @{thm conj2_False};
 | |
| 45 | val conj_True = @{thm conj_True};
 | |
| 46 | val conj_cong = @{thm conj_cong};
 | |
| 32921 | 47 | |
| 69597 | 48 | fun isFalse (Const (\<^const_name>\<open>False\<close>, _)) = true | 
| 25171 | 49 | | isFalse _ = false; | 
| 45363 | 50 | |
| 69597 | 51 | fun isTrue (Const (\<^const_name>\<open>True\<close>, _)) = true | 
| 25171 | 52 | | isTrue _ = false; | 
| 53 | ||
| 54 | in | |
| 32921 | 55 | |
| 25171 | 56 | val lazy_conj_simproc = | 
| 69597 | 57 | Simplifier.make_simproc \<^context> "lazy_conj_simp" | 
| 58 |    {lhss = [\<^term>\<open>P & Q\<close>],
 | |
| 61144 | 59 | proc = fn _ => fn ctxt => fn ct => | 
| 60 | (case Thm.term_of ct of | |
| 69597 | 61 | Const (\<^const_name>\<open>HOL.conj\<close>,_) $ P $ Q => | 
| 61144 | 62 | let | 
| 63 | val P_P' = Simplifier.rewrite ctxt (Thm.cterm_of ctxt P); | |
| 64 | val P' = P_P' |> Thm.prop_of |> Logic.dest_equals |> #2; | |
| 65 | in | |
| 66 | if isFalse P' then SOME (conj1_False OF [P_P']) | |
| 67 | else | |
| 68 | let | |
| 69 | val Q_Q' = Simplifier.rewrite ctxt (Thm.cterm_of ctxt Q); | |
| 70 | val Q' = Q_Q' |> Thm.prop_of |> Logic.dest_equals |> #2; | |
| 71 | in | |
| 72 | if isFalse Q' then SOME (conj2_False OF [Q_Q']) | |
| 73 | else if isTrue P' andalso isTrue Q' then SOME (conj_True OF [P_P', Q_Q']) | |
| 74 | else if P aconv P' andalso Q aconv Q' then NONE | |
| 75 | else SOME (conj_cong OF [P_P', Q_Q']) | |
| 76 | end | |
| 77 | end | |
| 62913 | 78 | | _ => NONE)}; | 
| 25171 | 79 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 80 | fun string_eq_simp_tac ctxt = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 81 | simp_tac (put_simpset HOL_basic_ss ctxt | 
| 68028 | 82 |     addsimps @{thms list.inject list.distinct char.inject
 | 
| 64630 
96015aecfeba
emphasize dedicated rewrite rules for congruences
 haftmann parents: 
62913diff
changeset | 83 | cong_exp_iff_simps simp_thms} | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 84 | addsimprocs [lazy_conj_simproc] | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 85 |     |> Simplifier.add_cong @{thm block_conj_cong});
 | 
| 32921 | 86 | |
| 25171 | 87 | end; | 
| 88 | ||
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 89 | val lookup_ss = | 
| 69597 | 90 | simpset_of (put_simpset HOL_basic_ss \<^context> | 
| 68028 | 91 |     addsimps (@{thms list.inject} @ @{thms char.inject}
 | 
| 58156 | 92 |       @ @{thms list.distinct} @ @{thms simp_thms}
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 93 |       @ [@{thm StateFun.lookup_update_id_same}, @{thm StateFun.id_id_cancel},
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 94 |         @{thm StateFun.lookup_update_same}, @{thm StateFun.lookup_update_other}])
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 95 | addsimprocs [lazy_conj_simproc] | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 96 | addSolver StateSpace.distinctNameSolver | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 97 |     |> fold Simplifier.add_cong @{thms block_conj_cong});
 | 
| 25171 | 98 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 99 | val ex_lookup_ss = | 
| 69597 | 100 |   simpset_of (put_simpset HOL_ss \<^context> addsimps @{thms StateFun.ex_id});
 | 
| 25171 | 101 | |
| 33519 | 102 | |
| 45363 | 103 | structure Data = Generic_Data | 
| 33519 | 104 | ( | 
| 45363 | 105 | type T = simpset * simpset * bool; (*lookup simpset, ex_lookup simpset, are simprocs installed*) | 
| 25171 | 106 | val empty = (empty_ss, empty_ss, false); | 
| 33519 | 107 | fun merge ((ss1, ex_ss1, b1), (ss2, ex_ss2, b2)) = | 
| 45363 | 108 | (merge_ss (ss1, ss2), merge_ss (ex_ss1, ex_ss2), b1 orelse b2); | 
| 33519 | 109 | ); | 
| 25171 | 110 | |
| 74586 
5ac762b53119
generalized component lookup for syntax and distinctness proofs. added some tracing.
 Norbert Schirmer <nschirmer@apple.com> parents: 
74561diff
changeset | 111 | fun trace_data context = | 
| 
5ac762b53119
generalized component lookup for syntax and distinctness proofs. added some tracing.
 Norbert Schirmer <nschirmer@apple.com> parents: 
74561diff
changeset | 112 | let | 
| 
5ac762b53119
generalized component lookup for syntax and distinctness proofs. added some tracing.
 Norbert Schirmer <nschirmer@apple.com> parents: 
74561diff
changeset | 113 | val (lookup_ss, ex_ss, active) = Data.get context | 
| 
5ac762b53119
generalized component lookup for syntax and distinctness proofs. added some tracing.
 Norbert Schirmer <nschirmer@apple.com> parents: 
74561diff
changeset | 114 | val pretty_ex_ss = Simplifier.pretty_simpset true (put_simpset ex_ss (Context.proof_of context)) | 
| 
5ac762b53119
generalized component lookup for syntax and distinctness proofs. added some tracing.
 Norbert Schirmer <nschirmer@apple.com> parents: 
74561diff
changeset | 115 | val pretty_lookup_ss = Simplifier.pretty_simpset true (put_simpset lookup_ss (Context.proof_of context)) | 
| 
5ac762b53119
generalized component lookup for syntax and distinctness proofs. added some tracing.
 Norbert Schirmer <nschirmer@apple.com> parents: 
74561diff
changeset | 116 | in | 
| 
5ac762b53119
generalized component lookup for syntax and distinctness proofs. added some tracing.
 Norbert Schirmer <nschirmer@apple.com> parents: 
74561diff
changeset | 117 |     tracing ("state_fun ex_ss: " ^ Pretty.string_of pretty_ex_ss ^ 
 | 
| 
5ac762b53119
generalized component lookup for syntax and distinctness proofs. added some tracing.
 Norbert Schirmer <nschirmer@apple.com> parents: 
74561diff
changeset | 118 | "\n ================= \n lookup_ss: " ^ Pretty.string_of pretty_lookup_ss ^ "\n " ^ | 
| 
5ac762b53119
generalized component lookup for syntax and distinctness proofs. added some tracing.
 Norbert Schirmer <nschirmer@apple.com> parents: 
74561diff
changeset | 119 |      "active: " ^ @{make_string} active)
 | 
| 
5ac762b53119
generalized component lookup for syntax and distinctness proofs. added some tracing.
 Norbert Schirmer <nschirmer@apple.com> parents: 
74561diff
changeset | 120 | end | 
| 
5ac762b53119
generalized component lookup for syntax and distinctness proofs. added some tracing.
 Norbert Schirmer <nschirmer@apple.com> parents: 
74561diff
changeset | 121 | |
| 58825 | 122 | val _ = Theory.setup (Context.theory_map (Data.put (lookup_ss, ex_lookup_ss, false))); | 
| 25171 | 123 | |
| 124 | val lookup_simproc = | |
| 69597 | 125 | Simplifier.make_simproc \<^context> "lookup_simp" | 
| 126 |    {lhss = [\<^term>\<open>lookup d n (update d' c m v s)\<close>],
 | |
| 61144 | 127 | proc = fn _ => fn ctxt => fn ct => | 
| 69597 | 128 | (case Thm.term_of ct of (Const (\<^const_name>\<open>StateFun.lookup\<close>, lT) $ destr $ n $ | 
| 129 | (s as Const (\<^const_name>\<open>StateFun.update\<close>, uT) $ _ $ _ $ _ $ _ $ _)) => | |
| 25171 | 130 | (let | 
| 131 | val (_::_::_::_::sT::_) = binder_types uT; | |
| 61144 | 132 | val mi = Term.maxidx_of_term (Thm.term_of ct); | 
| 69597 | 133 | fun mk_upds (Const (\<^const_name>\<open>StateFun.update\<close>, uT) $ d' $ c $ m $ v $ s) = | 
| 45363 | 134 | let | 
| 135 | val (_ :: _ :: _ :: fT :: _ :: _) = binder_types uT; | |
| 136 | val vT = domain_type fT; | |
| 137 | val (s', cnt) = mk_upds s; | |
| 138 | val (v', cnt') = | |
| 139 | (case v of | |
| 69597 | 140 | Const (\<^const_name>\<open>K_statefun\<close>, KT) $ v'' => | 
| 45363 | 141 | (case v'' of | 
| 69597 | 142 | (Const (\<^const_name>\<open>StateFun.lookup\<close>, _) $ | 
| 143 | (d as (Const (\<^const_name>\<open>Fun.id\<close>, _))) $ n' $ _) => | |
| 45363 | 144 | if d aconv c andalso n aconv m andalso m aconv n' | 
| 145 | then (v,cnt) (* Keep value so that | |
| 146 | lookup_update_id_same can fire *) | |
| 147 | else | |
| 69597 | 148 | (Const (\<^const_name>\<open>StateFun.K_statefun\<close>, KT) $ | 
| 45363 | 149 |                                   Var (("v", cnt), vT), cnt + 1)
 | 
| 150 | | _ => | |
| 69597 | 151 | (Const (\<^const_name>\<open>StateFun.K_statefun\<close>, KT) $ | 
| 45363 | 152 |                             Var (("v", cnt), vT), cnt + 1))
 | 
| 153 | | _ => (v, cnt)); | |
| 69597 | 154 | in (Const (\<^const_name>\<open>StateFun.update\<close>, uT) $ d' $ c $ m $ v' $ s', cnt') end | 
| 45363 | 155 |             | mk_upds s = (Var (("s", mi + 1), sT), mi + 2);
 | 
| 156 | ||
| 157 | val ct = | |
| 59643 | 158 | Thm.cterm_of ctxt | 
| 69597 | 159 | (Const (\<^const_name>\<open>StateFun.lookup\<close>, lT) $ destr $ n $ fst (mk_upds s)); | 
| 45363 | 160 | val basic_ss = #1 (Data.get (Context.Proof ctxt)); | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 161 | val ctxt' = ctxt |> Config.put simp_depth_limit 100 |> put_simpset basic_ss; | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 162 | val thm = Simplifier.rewrite ctxt' ct; | 
| 45363 | 163 | in | 
| 59582 | 164 | if (op aconv) (Logic.dest_equals (Thm.prop_of thm)) | 
| 45363 | 165 | then NONE | 
| 166 | else SOME thm | |
| 25171 | 167 | end | 
| 45361 | 168 | handle Option.Option => NONE) | 
| 62913 | 169 | | _ => NONE)}; | 
| 25171 | 170 | |
| 171 | ||
| 172 | local | |
| 45363 | 173 | |
| 32921 | 174 | val meta_ext = @{thm StateFun.meta_ext};
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 175 | val ss' = | 
| 69597 | 176 | simpset_of (put_simpset HOL_ss \<^context> addsimps | 
| 68028 | 177 |     (@{thm StateFun.update_apply} :: @{thm Fun.o_apply} :: @{thms list.inject} @ @{thms char.inject}
 | 
| 58156 | 178 |       @ @{thms list.distinct})
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 179 | addsimprocs [lazy_conj_simproc, StateSpace.distinct_simproc] | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 180 |     |> fold Simplifier.add_cong @{thms block_conj_cong});
 | 
| 45363 | 181 | |
| 25171 | 182 | in | 
| 45363 | 183 | |
| 25171 | 184 | val update_simproc = | 
| 69597 | 185 | Simplifier.make_simproc \<^context> "update_simp" | 
| 186 |    {lhss = [\<^term>\<open>update d c n v s\<close>],
 | |
| 61144 | 187 | proc = fn _ => fn ctxt => fn ct => | 
| 188 | (case Thm.term_of ct of | |
| 69597 | 189 | Const (\<^const_name>\<open>StateFun.update\<close>, uT) $ _ $ _ $ _ $ _ $ _ => | 
| 45363 | 190 | let | 
| 191 | val (_ :: _ :: _ :: _ :: sT :: _) = binder_types uT; | |
| 192 |               (*"('v => 'a1) => ('a2 => 'v) => 'n => ('a1 => 'a2) => ('n => 'v) => ('n => 'v)"*)
 | |
| 193 |             fun init_seed s = (Bound 0, Bound 0, [("s", sT)], [], false);
 | |
| 25171 | 194 | |
| 45363 | 195 | fun mk_comp f fT g gT = | 
| 196 | let val T = domain_type fT --> range_type gT | |
| 69597 | 197 | in (Const (\<^const_name>\<open>Fun.comp\<close>, gT --> fT --> T) $ g $ f, T) end; | 
| 25171 | 198 | |
| 45363 | 199 | fun mk_comps fs = foldl1 (fn ((f, fT), (g, gT)) => mk_comp g gT f fT) fs; | 
| 200 | ||
| 201 | fun append n c cT f fT d dT comps = | |
| 202 | (case AList.lookup (op aconv) comps n of | |
| 203 | SOME gTs => AList.update (op aconv) (n, [(c, cT), (f, fT), (d, dT)] @ gTs) comps | |
| 204 | | NONE => AList.update (op aconv) (n, [(c, cT), (f, fT), (d, dT)]) comps); | |
| 25171 | 205 | |
| 45363 | 206 | fun split_list (x :: xs) = let val (xs', y) = split_last xs in (x, xs', y) end | 
| 207 | | split_list _ = error "StateFun.split_list"; | |
| 25171 | 208 | |
| 45363 | 209 | fun merge_upds n comps = | 
| 210 | let val ((c, cT), fs, (d, dT)) = split_list (the (AList.lookup (op aconv) comps n)) | |
| 211 | in ((c, cT), fst (mk_comps fs), (d, dT)) end; | |
| 25171 | 212 | |
| 45363 | 213 | (* mk_updterm returns | 
| 214 | * - (orig-term-skeleton,simplified-term-skeleton, vars, b) | |
| 215 | * where boolean b tells if a simplification has occurred. | |
| 216 | "orig-term-skeleton = simplified-term-skeleton" is | |
| 217 | * the desired simplification rule. | |
| 218 | * The algorithm first walks down the updates to the seed-state while | |
| 219 | * memorising the updates in the already-table. While walking up the | |
| 220 | * updates again, the optimised term is constructed. | |
| 221 | *) | |
| 222 | fun mk_updterm already | |
| 69597 | 223 | ((upd as Const (\<^const_name>\<open>StateFun.update\<close>, uT)) $ d $ c $ n $ v $ s) = | 
| 45363 | 224 | let | 
| 225 | fun rest already = mk_updterm already; | |
| 226 | val (dT :: cT :: nT :: vT :: sT :: _) = binder_types uT; | |
| 227 |                       (*"('v => 'a1) => ('a2 => 'v) => 'n => ('a1 => 'a2) =>
 | |
| 228 |                             ('n => 'v) => ('n => 'v)"*)
 | |
| 229 | in | |
| 230 | if member (op aconv) already n then | |
| 231 | (case rest already s of | |
| 232 | (trm, trm', vars, comps, _) => | |
| 233 | let | |
| 234 | val i = length vars; | |
| 235 | val kv = (mk_name i n, vT); | |
| 236 | val kb = Bound i; | |
| 237 | val comps' = append n c cT kb vT d dT comps; | |
| 238 | in (upd $ d $ c $ n $ kb $ trm, trm', kv :: vars, comps',true) end) | |
| 239 | else | |
| 240 | (case rest (n :: already) s of | |
| 241 | (trm, trm', vars, comps, b) => | |
| 242 | let | |
| 243 | val i = length vars; | |
| 244 | val kv = (mk_name i n, vT); | |
| 245 | val kb = Bound i; | |
| 246 | val comps' = append n c cT kb vT d dT comps; | |
| 247 | val ((c', c'T), f', (d', d'T)) = merge_upds n comps'; | |
| 248 | val vT' = range_type d'T --> domain_type c'T; | |
| 249 | val upd' = | |
| 69597 | 250 | Const (\<^const_name>\<open>StateFun.update\<close>, | 
| 45363 | 251 | d'T --> c'T --> nT --> vT' --> sT --> sT); | 
| 252 | in | |
| 55972 | 253 | (upd $ d $ c $ n $ kb $ trm, upd' $ d' $ c' $ n $ f' $ trm', kv :: vars, | 
| 254 | comps', b) | |
| 45363 | 255 | end) | 
| 256 | end | |
| 257 | | mk_updterm _ t = init_seed t; | |
| 25171 | 258 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 259 | val ctxt0 = Config.put simp_depth_limit 100 ctxt; | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 260 | val ctxt1 = put_simpset ss' ctxt0; | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 261 | val ctxt2 = put_simpset (#1 (Data.get (Context.Proof ctxt0))) ctxt0; | 
| 45363 | 262 | in | 
| 61144 | 263 | (case mk_updterm [] (Thm.term_of ct) of | 
| 45363 | 264 | (trm, trm', vars, _, true) => | 
| 265 | let | |
| 266 | val eq1 = | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 267 | Goal.prove ctxt0 [] [] | 
| 46218 
ecf6375e2abb
renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
 wenzelm parents: 
45740diff
changeset | 268 | (Logic.list_all (vars, Logic.mk_equals (trm, trm'))) | 
| 60754 | 269 | (fn _ => resolve_tac ctxt0 [meta_ext] 1 THEN simp_tac ctxt1 1); | 
| 59582 | 270 | val eq2 = Simplifier.asm_full_rewrite ctxt2 (Thm.dest_equals_rhs (Thm.cprop_of eq1)); | 
| 45363 | 271 | in SOME (Thm.transitive eq1 eq2) end | 
| 272 | | _ => NONE) | |
| 273 | end | |
| 62913 | 274 | | _ => NONE)}; | 
| 25171 | 275 | |
| 45363 | 276 | end; | 
| 25171 | 277 | |
| 278 | ||
| 279 | local | |
| 45363 | 280 | |
| 39159 | 281 | val swap_ex_eq = @{thm StateFun.swap_ex_eq};
 | 
| 45363 | 282 | |
| 25171 | 283 | fun is_selector thy T sel = | 
| 45363 | 284 | let val (flds, more) = Record.get_recT_fields thy T | 
| 285 | in member (fn (s, (n, _)) => n = s) (more :: flds) sel end; | |
| 286 | ||
| 25171 | 287 | in | 
| 45363 | 288 | |
| 25171 | 289 | val ex_lookup_eq_simproc = | 
| 69597 | 290 | Simplifier.make_simproc \<^context> "ex_lookup_eq_simproc" | 
| 291 |    {lhss = [\<^term>\<open>Ex t\<close>],
 | |
| 61144 | 292 | proc = fn _ => fn ctxt => fn ct => | 
| 45363 | 293 | let | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 294 | val thy = Proof_Context.theory_of ctxt; | 
| 61144 | 295 | val t = Thm.term_of ct; | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 296 | |
| 45363 | 297 | val ex_lookup_ss = #2 (Data.get (Context.Proof ctxt)); | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 298 | val ctxt' = ctxt |> Config.put simp_depth_limit 100 |> put_simpset ex_lookup_ss; | 
| 45363 | 299 | fun prove prop = | 
| 300 | Goal.prove_global thy [] [] prop | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46218diff
changeset | 301 | (fn _ => Record.split_simp_tac ctxt [] (K ~1) 1 THEN simp_tac ctxt' 1); | 
| 25171 | 302 | |
| 45363 | 303 | fun mkeq (swap, Teq, lT, lo, d, n, x, s) i = | 
| 304 | let | |
| 305 | val (_ :: nT :: _) = binder_types lT; | |
| 306 |             (*  ('v => 'a) => 'n => ('n => 'v) => 'a *)
 | |
| 307 |             val x' = if not (Term.is_dependent x) then Bound 1 else raise TERM ("", [x]);
 | |
| 308 |             val n' = if not (Term.is_dependent n) then Bound 2 else raise TERM ("", [n]);
 | |
| 309 | val sel' = lo $ d $ n' $ s; | |
| 69597 | 310 | in (Const (\<^const_name>\<open>HOL.eq\<close>, Teq) $ sel' $ x', hd (binder_types Teq), nT, swap) end; | 
| 45363 | 311 | |
| 312 | fun dest_state (s as Bound 0) = s | |
| 313 | | dest_state (s as (Const (sel, sT) $ Bound 0)) = | |
| 314 | if is_selector thy (domain_type sT) sel then s | |
| 315 |               else raise TERM ("StateFun.ex_lookup_eq_simproc: not a record slector", [s])
 | |
| 316 |           | dest_state s = raise TERM ("StateFun.ex_lookup_eq_simproc: not a record slector", [s]);
 | |
| 25171 | 317 | |
| 45363 | 318 | fun dest_sel_eq | 
| 69597 | 319 | (Const (\<^const_name>\<open>HOL.eq\<close>, Teq) $ | 
| 320 | ((lo as (Const (\<^const_name>\<open>StateFun.lookup\<close>, lT))) $ d $ n $ s) $ X) = | |
| 45363 | 321 | (false, Teq, lT, lo, d, n, X, dest_state s) | 
| 322 | | dest_sel_eq | |
| 69597 | 323 | (Const (\<^const_name>\<open>HOL.eq\<close>, Teq) $ X $ | 
| 324 | ((lo as (Const (\<^const_name>\<open>StateFun.lookup\<close>, lT))) $ d $ n $ s)) = | |
| 45363 | 325 | (true, Teq, lT, lo, d, n, X, dest_state s) | 
| 326 |           | dest_sel_eq _ = raise TERM ("", []);
 | |
| 327 | in | |
| 328 | (case t of | |
| 69597 | 329 | Const (\<^const_name>\<open>Ex\<close>, Tex) $ Abs (s, T, t) => | 
| 45363 | 330 | (let | 
| 331 | val (eq, eT, nT, swap) = mkeq (dest_sel_eq t) 0; | |
| 332 | val prop = | |
| 46218 
ecf6375e2abb
renamed Term.list_all to Logic.list_all, in accordance to HOLogic.list_all;
 wenzelm parents: 
45740diff
changeset | 333 |                 Logic.list_all ([("n", nT), ("x", eT)],
 | 
| 69597 | 334 | Logic.mk_equals (Const (\<^const_name>\<open>Ex\<close>, Tex) $ Abs (s, T, eq), \<^term>\<open>True\<close>)); | 
| 45363 | 335 | val thm = Drule.export_without_context (prove prop); | 
| 336 | val thm' = if swap then swap_ex_eq OF [thm] else thm | |
| 337 | in SOME thm' end handle TERM _ => NONE) | |
| 338 | | _ => NONE) | |
| 62913 | 339 | end handle Option.Option => NONE}; | 
| 25171 | 340 | |
| 341 | end; | |
| 342 | ||
| 343 | val val_sfx = "V"; | |
| 344 | val val_prfx = "StateFun." | |
| 345 | fun deco base_prfx s = val_prfx ^ (base_prfx ^ suffix val_sfx s); | |
| 346 | ||
| 45363 | 347 | fun mkUpper str = | 
| 25171 | 348 | (case String.explode str of | 
| 349 | [] => "" | |
| 45363 | 350 | | c::cs => String.implode (Char.toUpper c :: cs)); | 
| 25171 | 351 | |
| 32952 | 352 | fun mkName (Type (T,args)) = implode (map mkName args) ^ mkUpper (Long_Name.base_name T) | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30289diff
changeset | 353 | | mkName (TFree (x,_)) = mkUpper (Long_Name.base_name x) | 
| 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30289diff
changeset | 354 | | mkName (TVar ((x,_),_)) = mkUpper (Long_Name.base_name x); | 
| 25171 | 355 | |
| 58354 
04ac60da613e
support (finite values of) codatatypes in Quickcheck
 blanchet parents: 
58156diff
changeset | 356 | fun is_datatype thy = is_some o BNF_LFP_Compat.get_info thy [BNF_LFP_Compat.Keep_Nesting]; | 
| 25171 | 357 | |
| 69597 | 358 | fun mk_map \<^type_name>\<open>List.list\<close> = Syntax.const \<^const_name>\<open>List.map\<close> | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30289diff
changeset | 359 |   | mk_map n = Syntax.const ("StateFun.map_" ^ Long_Name.base_name n);
 | 
| 25171 | 360 | |
| 45363 | 361 | fun gen_constr_destr comp prfx thy (Type (T, [])) = | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30289diff
changeset | 362 | Syntax.const (deco prfx (mkUpper (Long_Name.base_name T))) | 
| 25171 | 363 |   | gen_constr_destr comp prfx thy (T as Type ("fun",_)) =
 | 
| 45363 | 364 | let val (argTs, rangeT) = strip_type T; | 
| 365 | in | |
| 366 | comp | |
| 32952 | 367 | (Syntax.const (deco prfx (implode (map mkName argTs) ^ "Fun"))) | 
| 45363 | 368 | (fold (fn x => fn y => x $ y) | 
| 369 | (replicate (length argTs) (Syntax.const "StateFun.map_fun")) | |
| 370 | (gen_constr_destr comp prfx thy rangeT)) | |
| 371 | end | |
| 372 | | gen_constr_destr comp prfx thy (T' as Type (T, argTs)) = | |
| 373 | if is_datatype thy T | |
| 374 | then (* datatype args are recursively embedded into val *) | |
| 375 | (case argTs of | |
| 376 | [argT] => | |
| 377 | comp | |
| 378 | ((Syntax.const (deco prfx (mkUpper (Long_Name.base_name T))))) | |
| 379 | ((mk_map T $ gen_constr_destr comp prfx thy argT)) | |
| 380 |         | _ => raise (TYPE ("StateFun.gen_constr_destr", [T'], [])))
 | |
| 381 | else (* type args are not recursively embedded into val *) | |
| 382 | Syntax.const (deco prfx (implode (map mkName argTs) ^ mkUpper (Long_Name.base_name T))) | |
| 383 |   | gen_constr_destr thy _ _ T = raise (TYPE ("StateFun.gen_constr_destr", [T], []));
 | |
| 25171 | 384 | |
| 69597 | 385 | val mk_constr = gen_constr_destr (fn a => fn b => Syntax.const \<^const_name>\<open>Fun.comp\<close> $ a $ b) ""; | 
| 386 | val mk_destr = gen_constr_destr (fn a => fn b => Syntax.const \<^const_name>\<open>Fun.comp\<close> $ b $ a) "the_"; | |
| 45363 | 387 | |
| 58825 | 388 | val _ = | 
| 389 | Theory.setup | |
| 69597 | 390 | (Attrib.setup \<^binding>\<open>statefun_simp\<close> | 
| 58825 | 391 | (Scan.succeed (Thm.declaration_attribute (fn thm => fn context => | 
| 392 | let | |
| 393 | val ctxt = Context.proof_of context; | |
| 394 | val (lookup_ss, ex_lookup_ss, simprocs_active) = Data.get context; | |
| 395 | val (lookup_ss', ex_lookup_ss') = | |
| 59582 | 396 | (case Thm.concl_of thm of | 
| 69597 | 397 | (_ $ ((Const (\<^const_name>\<open>Ex\<close>, _) $ _))) => | 
| 58825 | 398 | (lookup_ss, simpset_map ctxt (Simplifier.add_simp thm) ex_lookup_ss) | 
| 399 | | _ => | |
| 400 | (simpset_map ctxt (Simplifier.add_simp thm) lookup_ss, ex_lookup_ss)); | |
| 401 | val activate_simprocs = | |
| 402 | if simprocs_active then I | |
| 403 | else Simplifier.map_ss (fn ctxt => ctxt addsimprocs [lookup_simproc, update_simproc]); | |
| 404 | in | |
| 405 | context | |
| 406 | |> activate_simprocs | |
| 407 | |> Data.put (lookup_ss', ex_lookup_ss', true) | |
| 408 | end))) | |
| 409 | "simplification in statespaces"); | |
| 45363 | 410 | |
| 411 | end; |