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