src/HOL/HOLCF/Tools/Domain/domain_isomorphism.ML
author wenzelm
Sat, 14 Dec 2013 17:28:05 +0100
changeset 54742 7a86358a3c0b
parent 52788 da1fdbfebd39
child 54895 515630483010
permissions -rw-r--r--
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.; clarified tool context in some boundary cases;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 41430
diff changeset
     1
(*  Title:      HOL/HOLCF/Tools/Domain/domain_isomorphism.ML
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
     2
    Author:     Brian Huffman
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
     3
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
     4
Defines new types satisfying the given domain equations.
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
     5
*)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
     6
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
     7
signature DOMAIN_ISOMORPHISM =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
     8
sig
35489
dd02201d95b6 add function define_take_functions
huffman
parents: 35480
diff changeset
     9
  val domain_isomorphism :
35657
0537c34c6067 pass take_induct_info as an argument to comp_theorems
huffman
parents: 35654
diff changeset
    10
      (string list * binding * mixfix * typ
0537c34c6067 pass take_induct_info as an argument to comp_theorems
huffman
parents: 35654
diff changeset
    11
       * (binding * binding) option) list ->
0537c34c6067 pass take_induct_info as an argument to comp_theorems
huffman
parents: 35654
diff changeset
    12
      theory ->
0537c34c6067 pass take_induct_info as an argument to comp_theorems
huffman
parents: 35654
diff changeset
    13
      (Domain_Take_Proofs.iso_info list
0537c34c6067 pass take_induct_info as an argument to comp_theorems
huffman
parents: 35654
diff changeset
    14
       * Domain_Take_Proofs.take_induct_info) * theory
0537c34c6067 pass take_induct_info as an argument to comp_theorems
huffman
parents: 35654
diff changeset
    15
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
    16
  val define_map_functions :
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
    17
      (binding * Domain_Take_Proofs.iso_info) list ->
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
    18
      theory ->
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
    19
      {
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
    20
        map_consts : term list,
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
    21
        map_apply_thms : thm list,
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
    22
        map_unfold_thms : thm list,
41324
1383653efec3 make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents: 41292
diff changeset
    23
        map_cont_thm : thm,
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
    24
        deflation_map_thms : thm list
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
    25
      }
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
    26
      * theory
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
    27
35489
dd02201d95b6 add function define_take_functions
huffman
parents: 35480
diff changeset
    28
  val domain_isomorphism_cmd :
34149
a0efb4754cb7 add 'morphisms' option to domain_isomorphism command
huffman
parents: 33971
diff changeset
    29
    (string list * binding * mixfix * string * (binding * binding) option) list
a0efb4754cb7 add 'morphisms' option to domain_isomorphism command
huffman
parents: 33971
diff changeset
    30
      -> theory -> theory
40218
f7d4d023a899 make domain package work with non-cpo argument types
huffman
parents: 40216
diff changeset
    31
40216
366309dfaf60 use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents: 40037
diff changeset
    32
  val setup : theory -> theory
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    33
end
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    34
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents: 35498
diff changeset
    35
structure Domain_Isomorphism : DOMAIN_ISOMORPHISM =
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    36
struct
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    37
40833
4f130bd9e17e internal domain package proofs use cont2cont simproc instead of a fixed list of cont rules
huffman
parents: 40832
diff changeset
    38
val beta_ss =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
    39
  simpset_of (put_simpset HOL_basic_ss @{context}
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
    40
    addsimps @{thms simp_thms} addsimprocs [@{simproc beta_cfun_proc}])
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
    41
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    42
fun is_cpo thy T = Sign.of_sort thy (T, @{sort cpo})
40218
f7d4d023a899 make domain package work with non-cpo argument types
huffman
parents: 40216
diff changeset
    43
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    44
(******************************************************************************)
33791
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    45
(******************************** theory data *********************************)
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    46
(******************************************************************************)
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    47
40216
366309dfaf60 use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents: 40037
diff changeset
    48
structure RepData = Named_Thms
33791
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    49
(
45294
3c5d3d286055 tuned Named_Thms: proper binding;
wenzelm
parents: 44080
diff changeset
    50
  val name = @{binding domain_defl_simps}
40216
366309dfaf60 use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents: 40037
diff changeset
    51
  val description = "theorems like DEFL('a t) = t_defl$DEFL('a)"
366309dfaf60 use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents: 40037
diff changeset
    52
)
366309dfaf60 use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents: 40037
diff changeset
    53
366309dfaf60 use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents: 40037
diff changeset
    54
structure IsodeflData = Named_Thms
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents: 35498
diff changeset
    55
(
45294
3c5d3d286055 tuned Named_Thms: proper binding;
wenzelm
parents: 44080
diff changeset
    56
  val name = @{binding domain_isodefl}
40216
366309dfaf60 use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents: 40037
diff changeset
    57
  val description = "theorems like isodefl d t ==> isodefl (foo_map$d) (foo_defl$t)"
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    58
)
33809
033831fd9ef3 store map_ID thms in theory data; automate proofs of reach lemmas
huffman
parents: 33807
diff changeset
    59
40489
0f37a553bc1a implement map_of_typ using Pattern.rewrite_term
huffman
parents: 40487
diff changeset
    60
val setup = RepData.setup #> IsodeflData.setup
33801
e8535acd302c copy_of_dtyp uses map table from theory data
huffman
parents: 33797
diff changeset
    61
33791
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    62
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
    63
(******************************************************************************)
35475
979019ab92eb move common functions into new file holcf_library.ML
huffman
parents: 35474
diff changeset
    64
(************************** building types and terms **************************)
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    65
(******************************************************************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    66
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    67
open HOLCF_Library
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    68
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    69
infixr 6 ->>
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    70
infixr -->>
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    71
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    72
val udomT = @{typ udom}
41287
029a6fc1bfb8 type 'defl' takes a type parameter again (cf. b525988432e9)
huffman
parents: 40833
diff changeset
    73
val deflT = @{typ "udom defl"}
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41287
diff changeset
    74
val udeflT = @{typ "udom u defl"}
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    75
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39974
diff changeset
    76
fun mk_DEFL T =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    77
  Const (@{const_name defl}, Term.itselfT T --> deflT) $ Logic.mk_type T
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
    78
40487
1320a0747974 implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents: 40486
diff changeset
    79
fun dest_DEFL (Const (@{const_name defl}, _) $ t) = Logic.dest_type t
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    80
  | dest_DEFL t = raise TERM ("dest_DEFL", [t])
40487
1320a0747974 implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents: 40486
diff changeset
    81
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
    82
fun mk_LIFTDEFL T =
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41287
diff changeset
    83
  Const (@{const_name liftdefl}, Term.itselfT T --> udeflT) $ Logic.mk_type T
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
    84
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
    85
fun dest_LIFTDEFL (Const (@{const_name liftdefl}, _) $ t) = Logic.dest_type t
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    86
  | dest_LIFTDEFL t = raise TERM ("dest_LIFTDEFL", [t])
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
    87
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    88
fun mk_u_defl t = mk_capply (@{const "u_defl"}, t)
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
    89
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    90
fun emb_const T = Const (@{const_name emb}, T ->> udomT)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    91
fun prj_const T = Const (@{const_name prj}, udomT ->> T)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    92
fun coerce_const (T, U) = mk_cfcomp (prj_const U, emb_const T)
33778
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
    93
33788
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
    94
fun isodefl_const T =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    95
  Const (@{const_name isodefl}, (T ->> T) --> deflT --> HOLogic.boolT)
33788
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
    96
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41287
diff changeset
    97
fun isodefl'_const T =
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41287
diff changeset
    98
  Const (@{const_name isodefl'}, (T ->> T) --> udeflT --> HOLogic.boolT)
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41287
diff changeset
    99
35480
7a1f285cad25 domain_isomorphism package proves deflation rules for map functions
huffman
parents: 35479
diff changeset
   100
fun mk_deflation t =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   101
  Const (@{const_name deflation}, Term.fastype_of t --> boolT) $ t
35480
7a1f285cad25 domain_isomorphism package proves deflation rules for map functions
huffman
parents: 35479
diff changeset
   102
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   103
(* splits a cterm into the right and lefthand sides of equality *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   104
fun dest_eqs t = HOLogic.dest_eq (HOLogic.dest_Trueprop t)
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   105
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   106
fun mk_eqs (t, u) = HOLogic.mk_Trueprop (HOLogic.mk_eq (t, u))
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   107
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   108
(******************************************************************************)
35489
dd02201d95b6 add function define_take_functions
huffman
parents: 35480
diff changeset
   109
(****************************** isomorphism info ******************************)
dd02201d95b6 add function define_take_functions
huffman
parents: 35480
diff changeset
   110
(******************************************************************************)
dd02201d95b6 add function define_take_functions
huffman
parents: 35480
diff changeset
   111
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents: 35498
diff changeset
   112
fun deflation_abs_rep (info : Domain_Take_Proofs.iso_info) : thm =
35489
dd02201d95b6 add function define_take_functions
huffman
parents: 35480
diff changeset
   113
  let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   114
    val abs_iso = #abs_inverse info
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   115
    val rep_iso = #rep_inverse info
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   116
    val thm = @{thm deflation_abs_rep} OF [abs_iso, rep_iso]
35489
dd02201d95b6 add function define_take_functions
huffman
parents: 35480
diff changeset
   117
  in
36241
2a4cec6bcae2 replace many uses of Drule.export_without_context with Drule.zero_var_indexes
huffman
parents: 35842
diff changeset
   118
    Drule.zero_var_indexes thm
35489
dd02201d95b6 add function define_take_functions
huffman
parents: 35480
diff changeset
   119
  end
dd02201d95b6 add function define_take_functions
huffman
parents: 35480
diff changeset
   120
dd02201d95b6 add function define_take_functions
huffman
parents: 35480
diff changeset
   121
(******************************************************************************)
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   122
(*************** fixed-point definitions and unfolding theorems ***************)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   123
(******************************************************************************)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   124
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 44066
diff changeset
   125
fun mk_projs []      _ = []
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   126
  | mk_projs (x::[]) t = [(x, t)]
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   127
  | mk_projs (x::xs) t = (x, mk_fst t) :: mk_projs xs (mk_snd t)
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   128
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   129
fun add_fixdefs
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   130
    (spec : (binding * term) list)
41324
1383653efec3 make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents: 41292
diff changeset
   131
    (thy : theory) : (thm list * thm list * thm) * theory =
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   132
  let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   133
    val binds = map fst spec
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   134
    val (lhss, rhss) = ListPair.unzip (map (dest_eqs o snd) spec)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   135
    val functional = lambda_tuple lhss (mk_tuple rhss)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   136
    val fixpoint = mk_fix (mk_cabs functional)
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   137
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   138
    (* project components of fixpoint *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   139
    val projs = mk_projs lhss fixpoint
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   140
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   141
    (* convert parameters to lambda abstractions *)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   142
    fun mk_eqn (lhs, rhs) =
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   143
        case lhs of
40327
1dfdbd66093a renamed {Rep,Abs}_CFun to {Rep,Abs}_cfun
huffman
parents: 40326
diff changeset
   144
          Const (@{const_name Rep_cfun}, _) $ f $ (x as Free _) =>
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   145
            mk_eqn (f, big_lambda x rhs)
40218
f7d4d023a899 make domain package work with non-cpo argument types
huffman
parents: 40216
diff changeset
   146
        | f $ Const (@{const_name TYPE}, T) =>
f7d4d023a899 make domain package work with non-cpo argument types
huffman
parents: 40216
diff changeset
   147
            mk_eqn (f, Abs ("t", T, rhs))
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   148
        | Const _ => Logic.mk_equals (lhs, rhs)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   149
        | _ => raise TERM ("lhs not of correct form", [lhs, rhs])
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   150
    val eqns = map mk_eqn projs
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   151
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   152
    (* register constant definitions *)
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   153
    val (fixdef_thms, thy) =
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37744
diff changeset
   154
      (Global_Theory.add_defs false o map Thm.no_attributes)
46909
3c73a121a387 more explicit indication of def names;
wenzelm
parents: 46708
diff changeset
   155
        (map Thm.def_binding binds ~~ eqns) thy
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   156
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   157
    (* prove applied version of definitions *)
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   158
    fun prove_proj (lhs, rhs) =
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   159
      let
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   160
        fun tac ctxt = rewrite_goals_tac ctxt fixdef_thms THEN
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   161
          (simp_tac (Simplifier.global_context thy beta_ss)) 1
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   162
        val goal = Logic.mk_equals (lhs, rhs)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   163
      in Goal.prove_global thy [] [] goal (tac o #context) end
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   164
    val proj_thms = map prove_proj projs
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   165
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   166
    (* mk_tuple lhss == fixpoint *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   167
    fun pair_equalI (thm1, thm2) = @{thm Pair_equalI} OF [thm1, thm2]
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   168
    val tuple_fixdef_thm = foldr1 pair_equalI proj_thms
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   169
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   170
    val cont_thm =
40833
4f130bd9e17e internal domain package proofs use cont2cont simproc instead of a fixed list of cont rules
huffman
parents: 40832
diff changeset
   171
      let
4f130bd9e17e internal domain package proofs use cont2cont simproc instead of a fixed list of cont rules
huffman
parents: 40832
diff changeset
   172
        val prop = mk_trp (mk_cont functional)
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42151
diff changeset
   173
        val rules = Cont2ContData.get (Proof_Context.init_global thy)
40833
4f130bd9e17e internal domain package proofs use cont2cont simproc instead of a fixed list of cont rules
huffman
parents: 40832
diff changeset
   174
        val tac = REPEAT_ALL_NEW (match_tac rules) 1
4f130bd9e17e internal domain package proofs use cont2cont simproc instead of a fixed list of cont rules
huffman
parents: 40832
diff changeset
   175
      in
4f130bd9e17e internal domain package proofs use cont2cont simproc instead of a fixed list of cont rules
huffman
parents: 40832
diff changeset
   176
        Goal.prove_global thy [] [] prop (K tac)
4f130bd9e17e internal domain package proofs use cont2cont simproc instead of a fixed list of cont rules
huffman
parents: 40832
diff changeset
   177
      end
4f130bd9e17e internal domain package proofs use cont2cont simproc instead of a fixed list of cont rules
huffman
parents: 40832
diff changeset
   178
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   179
    val tuple_unfold_thm =
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   180
      (@{thm def_cont_fix_eq} OF [tuple_fixdef_thm, cont_thm])
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42151
diff changeset
   181
      |> Local_Defs.unfold (Proof_Context.init_global thy) @{thms split_conv}
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   182
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 44066
diff changeset
   183
    fun mk_unfold_thms [] _ = []
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   184
      | mk_unfold_thms (n::[]) thm = [(n, thm)]
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   185
      | mk_unfold_thms (n::ns) thm = let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   186
          val thmL = thm RS @{thm Pair_eqD1}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   187
          val thmR = thm RS @{thm Pair_eqD2}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   188
        in (n, thmL) :: mk_unfold_thms ns thmR end
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   189
    val unfold_binds = map (Binding.suffix_name "_unfold") binds
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   190
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   191
    (* register unfold theorems *)
33777
69eae9bca167 get rid of numbers on thy variables
huffman
parents: 33776
diff changeset
   192
    val (unfold_thms, thy) =
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37744
diff changeset
   193
      (Global_Theory.add_thms o map (Thm.no_attributes o apsnd Drule.zero_var_indexes))
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   194
        (mk_unfold_thms unfold_binds tuple_unfold_thm) thy
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   195
  in
41324
1383653efec3 make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents: 41292
diff changeset
   196
    ((proj_thms, unfold_thms, cont_thm), thy)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   197
  end
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   198
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   199
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   200
(******************************************************************************)
33790
b2ff505e30f8 replace defl_tab and map_tab with theory data
huffman
parents: 33789
diff changeset
   201
(****************** deflation combinators and map functions *******************)
b2ff505e30f8 replace defl_tab and map_tab with theory data
huffman
parents: 33789
diff changeset
   202
(******************************************************************************)
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   203
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   204
fun defl_of_typ
40218
f7d4d023a899 make domain package work with non-cpo argument types
huffman
parents: 40216
diff changeset
   205
    (thy : theory)
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   206
    (tab1 : (typ * term) list)
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   207
    (tab2 : (typ * term) list)
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   208
    (T : typ) : term =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   209
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42151
diff changeset
   210
    val defl_simps = RepData.get (Proof_Context.init_global thy)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   211
    val rules = map (Thm.concl_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq) defl_simps
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   212
    val rules' = map (apfst mk_DEFL) tab1 @ map (apfst mk_LIFTDEFL) tab2
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   213
    fun proc1 t =
40487
1320a0747974 implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents: 40486
diff changeset
   214
      (case dest_DEFL t of
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   215
        TFree (a, _) => SOME (Free ("d" ^ Library.unprefix "'" a, deflT))
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   216
      | _ => NONE) handle TERM _ => NONE
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   217
    fun proc2 t =
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   218
      (case dest_LIFTDEFL t of
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41287
diff changeset
   219
        TFree (a, _) => SOME (Free ("p" ^ Library.unprefix "'" a, udeflT))
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   220
      | _ => NONE) handle TERM _ => NONE
40487
1320a0747974 implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents: 40486
diff changeset
   221
  in
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   222
    Pattern.rewrite_term thy (rules @ rules') [proc1, proc2] (mk_DEFL T)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   223
  end
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   224
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   225
(******************************************************************************)
35477
b972233773dd add function define_const
huffman
parents: 35475
diff changeset
   226
(********************* declaring definitions and theorems *********************)
b972233773dd add function define_const
huffman
parents: 35475
diff changeset
   227
(******************************************************************************)
b972233773dd add function define_const
huffman
parents: 35475
diff changeset
   228
b972233773dd add function define_const
huffman
parents: 35475
diff changeset
   229
fun define_const
b972233773dd add function define_const
huffman
parents: 35475
diff changeset
   230
    (bind : binding, rhs : term)
b972233773dd add function define_const
huffman
parents: 35475
diff changeset
   231
    (thy : theory)
b972233773dd add function define_const
huffman
parents: 35475
diff changeset
   232
    : (term * thm) * theory =
b972233773dd add function define_const
huffman
parents: 35475
diff changeset
   233
  let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   234
    val typ = Term.fastype_of rhs
42375
774df7c59508 report Name_Space.declare/define, relatively to context;
wenzelm
parents: 42361
diff changeset
   235
    val (const, thy) = Sign.declare_const_global ((bind, typ), NoSyn) thy
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   236
    val eqn = Logic.mk_equals (const, rhs)
46909
3c73a121a387 more explicit indication of def names;
wenzelm
parents: 46708
diff changeset
   237
    val def = Thm.no_attributes (Thm.def_binding bind, eqn)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   238
    val (def_thm, thy) = yield_singleton (Global_Theory.add_defs false) def thy
35477
b972233773dd add function define_const
huffman
parents: 35475
diff changeset
   239
  in
b972233773dd add function define_const
huffman
parents: 35475
diff changeset
   240
    ((const, def_thm), thy)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   241
  end
35477
b972233773dd add function define_const
huffman
parents: 35475
diff changeset
   242
35773
cae4f840d15d more consistent use of qualified bindings
huffman
parents: 35659
diff changeset
   243
fun add_qualified_thm name (dbind, thm) =
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37744
diff changeset
   244
    yield_singleton Global_Theory.add_thms
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   245
      ((Binding.qualified true name dbind, thm), [])
35489
dd02201d95b6 add function define_take_functions
huffman
parents: 35480
diff changeset
   246
dd02201d95b6 add function define_take_functions
huffman
parents: 35480
diff changeset
   247
(******************************************************************************)
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   248
(*************************** defining map functions ***************************)
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   249
(******************************************************************************)
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   250
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   251
fun define_map_functions
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   252
    (spec : (binding * Domain_Take_Proofs.iso_info) list)
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   253
    (thy : theory) =
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   254
  let
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   255
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   256
    (* retrieve components of spec *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   257
    val dbinds = map fst spec
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   258
    val iso_infos = map snd spec
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   259
    val dom_eqns = map (fn x => (#absT x, #repT x)) iso_infos
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   260
    val rep_abs_consts = map (fn x => (#rep_const x, #abs_const x)) iso_infos
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   261
40218
f7d4d023a899 make domain package work with non-cpo argument types
huffman
parents: 40216
diff changeset
   262
    fun mapT (T as Type (_, Ts)) =
f7d4d023a899 make domain package work with non-cpo argument types
huffman
parents: 40216
diff changeset
   263
        (map (fn T => T ->> T) (filter (is_cpo thy) Ts)) -->> (T ->> T)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   264
      | mapT T = T ->> T
40218
f7d4d023a899 make domain package work with non-cpo argument types
huffman
parents: 40216
diff changeset
   265
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   266
    (* declare map functions *)
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 44066
diff changeset
   267
    fun declare_map_const (tbind, (lhsT, _)) thy =
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   268
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   269
        val map_type = mapT lhsT
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   270
        val map_bind = Binding.suffix_name "_map" tbind
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   271
      in
42375
774df7c59508 report Name_Space.declare/define, relatively to context;
wenzelm
parents: 42361
diff changeset
   272
        Sign.declare_const_global ((map_bind, map_type), NoSyn) thy
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   273
      end
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   274
    val (map_consts, thy) = thy |>
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   275
      fold_map declare_map_const (dbinds ~~ dom_eqns)
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   276
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   277
    (* defining equations for map functions *)
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   278
    local
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   279
      fun unprime a = Library.unprefix "'" a
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   280
      fun mapvar T = Free (unprime (fst (dest_TFree T)), T ->> T)
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   281
      fun map_lhs (map_const, lhsT) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   282
          (lhsT, list_ccomb (map_const, map mapvar (filter (is_cpo thy) (snd (dest_Type lhsT)))))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   283
      val tab1 = map map_lhs (map_consts ~~ map fst dom_eqns)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   284
      val Ts = (snd o dest_Type o fst o hd) dom_eqns
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   285
      val tab = (Ts ~~ map mapvar Ts) @ tab1
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 44066
diff changeset
   286
      fun mk_map_spec (((rep_const, abs_const), _), (lhsT, rhsT)) =
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   287
        let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   288
          val lhs = Domain_Take_Proofs.map_of_typ thy tab lhsT
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   289
          val body = Domain_Take_Proofs.map_of_typ thy tab rhsT
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   290
          val rhs = mk_cfcomp (abs_const, mk_cfcomp (body, rep_const))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   291
        in mk_eqs (lhs, rhs) end
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   292
    in
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   293
      val map_specs =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   294
          map mk_map_spec (rep_abs_consts ~~ map_consts ~~ dom_eqns)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   295
    end
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   296
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   297
    (* register recursive definition of map functions *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   298
    val map_binds = map (Binding.suffix_name "_map") dbinds
41324
1383653efec3 make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents: 41292
diff changeset
   299
    val ((map_apply_thms, map_unfold_thms, map_cont_thm), thy) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   300
      add_fixdefs (map_binds ~~ map_specs) thy
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   301
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   302
    (* prove deflation theorems for map functions *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   303
    val deflation_abs_rep_thms = map deflation_abs_rep iso_infos
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   304
    val deflation_map_thm =
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   305
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   306
        fun unprime a = Library.unprefix "'" a
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   307
        fun mk_f T = Free (unprime (fst (dest_TFree T)), T ->> T)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   308
        fun mk_assm T = mk_trp (mk_deflation (mk_f T))
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 44066
diff changeset
   309
        fun mk_goal (map_const, (lhsT, _)) =
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   310
          let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   311
            val (_, Ts) = dest_Type lhsT
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   312
            val map_term = list_ccomb (map_const, map mk_f (filter (is_cpo thy) Ts))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   313
          in mk_deflation map_term end
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   314
        val assms = (map mk_assm o filter (is_cpo thy) o snd o dest_Type o fst o hd) dom_eqns
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   315
        val goals = map mk_goal (map_consts ~~ dom_eqns)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   316
        val goal = mk_trp (foldr1 HOLogic.mk_conj goals)
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   317
        val adm_rules =
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   318
          @{thms adm_conj adm_subst [OF _ adm_deflation]
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   319
                 cont2cont_fst cont2cont_snd cont_id}
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   320
        val bottom_rules =
41430
1aa23e9f2c87 change some lemma names containing 'UU' to 'bottom'
huffman
parents: 41324
diff changeset
   321
          @{thms fst_strict snd_strict deflation_bottom simp_thms}
41324
1383653efec3 make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents: 41292
diff changeset
   322
        val tuple_rules =
1383653efec3 make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents: 41292
diff changeset
   323
          @{thms split_def fst_conv snd_conv}
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   324
        val deflation_rules =
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   325
          @{thms conjI deflation_ID}
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   326
          @ deflation_abs_rep_thms
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   327
          @ Domain_Take_Proofs.get_deflation_thms thy
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   328
      in
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   329
        Goal.prove_global thy [] assms goal (fn {prems, context = ctxt} =>
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   330
         EVERY
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   331
          [rewrite_goals_tac ctxt map_apply_thms,
41324
1383653efec3 make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents: 41292
diff changeset
   332
           rtac (map_cont_thm RS @{thm cont_fix_ind}) 1,
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   333
           REPEAT (resolve_tac adm_rules 1),
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   334
           simp_tac (put_simpset HOL_basic_ss ctxt addsimps bottom_rules) 1,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   335
           simp_tac (put_simpset HOL_basic_ss ctxt addsimps tuple_rules) 1,
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   336
           REPEAT (etac @{thm conjE} 1),
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   337
           REPEAT (resolve_tac (deflation_rules @ prems) 1 ORELSE atac 1)])
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   338
      end
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 44066
diff changeset
   339
    fun conjuncts [] _ = []
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   340
      | conjuncts (n::[]) thm = [(n, thm)]
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   341
      | conjuncts (n::ns) thm = let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   342
          val thmL = thm RS @{thm conjunct1}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   343
          val thmR = thm RS @{thm conjunct2}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   344
        in (n, thmL):: conjuncts ns thmR end
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   345
    val deflation_map_binds = dbinds |>
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   346
        map (Binding.prefix_name "deflation_" o Binding.suffix_name "_map")
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   347
    val (deflation_map_thms, thy) = thy |>
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37744
diff changeset
   348
      (Global_Theory.add_thms o map (Thm.no_attributes o apsnd Drule.zero_var_indexes))
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   349
        (conjuncts deflation_map_binds deflation_map_thm)
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   350
40737
2037021f034f remove map function names from domain package theory data
huffman
parents: 40575
diff changeset
   351
    (* register indirect recursion in theory data *)
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   352
    local
40737
2037021f034f remove map function names from domain package theory data
huffman
parents: 40575
diff changeset
   353
      fun register_map (dname, args) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   354
        Domain_Take_Proofs.add_rec_type (dname, args)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   355
      val dnames = map (fst o dest_Type o fst) dom_eqns
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   356
      fun args (T, _) = case T of Type (_, Ts) => map (is_cpo thy) Ts | _ => []
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   357
      val argss = map args dom_eqns
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   358
    in
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   359
      val thy =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   360
          fold register_map (dnames ~~ argss) thy
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   361
    end
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   362
40216
366309dfaf60 use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents: 40037
diff changeset
   363
    (* register deflation theorems *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   364
    val thy = fold Domain_Take_Proofs.add_deflation_thm deflation_map_thms thy
40216
366309dfaf60 use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents: 40037
diff changeset
   365
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   366
    val result =
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   367
      {
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   368
        map_consts = map_consts,
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   369
        map_apply_thms = map_apply_thms,
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   370
        map_unfold_thms = map_unfold_thms,
41324
1383653efec3 make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents: 41292
diff changeset
   371
        map_cont_thm = map_cont_thm,
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   372
        deflation_map_thms = deflation_map_thms
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   373
      }
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   374
  in
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   375
    (result, thy)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   376
  end
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   377
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   378
(******************************************************************************)
35477
b972233773dd add function define_const
huffman
parents: 35475
diff changeset
   379
(******************************* main function ********************************)
b972233773dd add function define_const
huffman
parents: 35475
diff changeset
   380
(******************************************************************************)
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   381
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   382
fun read_typ thy str sorts =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   383
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42151
diff changeset
   384
    val ctxt = Proof_Context.init_global thy
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   385
      |> fold (Variable.declare_typ o TFree) sorts
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   386
    val T = Syntax.read_typ ctxt str
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   387
  in (T, Term.add_tfreesT T sorts) end
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   388
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   389
fun cert_typ sign raw_T sorts =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   390
  let
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   391
    val T = Type.no_tvars (Sign.certify_typ sign raw_T)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   392
      handle TYPE (msg, _, _) => error msg
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   393
    val sorts' = Term.add_tfreesT T sorts
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   394
    val _ =
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   395
      case duplicates (op =) (map fst sorts') of
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   396
        [] => ()
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   397
      | dups => error ("Inconsistent sort constraints for " ^ commas dups)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   398
  in (T, sorts') end
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   399
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   400
fun gen_domain_isomorphism
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   401
    (prep_typ: theory -> 'a -> (string * sort) list -> typ * (string * sort) list)
34149
a0efb4754cb7 add 'morphisms' option to domain_isomorphism command
huffman
parents: 33971
diff changeset
   402
    (doms_raw: (string list * binding * mixfix * 'a * (binding * binding) option) list)
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   403
    (thy: theory)
35657
0537c34c6067 pass take_induct_info as an argument to comp_theorems
huffman
parents: 35654
diff changeset
   404
    : (Domain_Take_Proofs.iso_info list
0537c34c6067 pass take_induct_info as an argument to comp_theorems
huffman
parents: 35654
diff changeset
   405
       * Domain_Take_Proofs.take_induct_info) * theory =
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   406
  let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   407
    val _ = Theory.requires thy "Domain" "domain isomorphisms"
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   408
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   409
    (* this theory is used just for parsing *)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   410
    val tmp_thy = thy |>
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 44066
diff changeset
   411
      Sign.add_types_global (map (fn (tvs, tbind, mx, _, _) =>
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   412
        (tbind, length tvs, mx)) doms_raw)
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   413
34149
a0efb4754cb7 add 'morphisms' option to domain_isomorphism command
huffman
parents: 33971
diff changeset
   414
    fun prep_dom thy (vs, t, mx, typ_raw, morphs) sorts =
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   415
      let val (typ, sorts') = prep_typ thy typ_raw sorts
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   416
      in ((vs, t, mx, typ, morphs), sorts') end
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   417
34149
a0efb4754cb7 add 'morphisms' option to domain_isomorphism command
huffman
parents: 33971
diff changeset
   418
    val (doms : (string list * binding * mixfix * typ * (binding * binding) option) list,
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   419
         sorts : (string * sort) list) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   420
      fold_map (prep_dom tmp_thy) doms_raw []
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   421
40486
0f2ae76c2e1d add function the_sort
huffman
parents: 40327
diff changeset
   422
    (* lookup function for sorts of type variables *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   423
    fun the_sort v = the (AList.lookup (op =) sorts v)
40486
0f2ae76c2e1d add function the_sort
huffman
parents: 40327
diff changeset
   424
40490
05be0c37db1d add bifiniteness check for domain_isomorphism command
huffman
parents: 40489
diff changeset
   425
    (* declare arities in temporary theory *)
40487
1320a0747974 implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents: 40486
diff changeset
   426
    val tmp_thy =
1320a0747974 implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents: 40486
diff changeset
   427
      let
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 44066
diff changeset
   428
        fun arity (vs, tbind, _, _, _) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   429
          (Sign.full_name thy tbind, map the_sort vs, @{sort "domain"})
40487
1320a0747974 implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents: 40486
diff changeset
   430
      in
51685
385ef6706252 more standard module name Axclass (according to file name);
wenzelm
parents: 49759
diff changeset
   431
        fold Axclass.axiomatize_arity (map arity doms) tmp_thy
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   432
      end
40487
1320a0747974 implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents: 40486
diff changeset
   433
40490
05be0c37db1d add bifiniteness check for domain_isomorphism command
huffman
parents: 40489
diff changeset
   434
    (* check bifiniteness of right-hand sides *)
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 44066
diff changeset
   435
    fun check_rhs (_, _, _, rhs, _) =
40497
d2e876d6da8c rename class 'bifinite' to 'domain'
huffman
parents: 40494
diff changeset
   436
      if Sign.of_sort tmp_thy (rhs, @{sort "domain"}) then ()
d2e876d6da8c rename class 'bifinite' to 'domain'
huffman
parents: 40494
diff changeset
   437
      else error ("Type not of sort domain: " ^
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   438
        quote (Syntax.string_of_typ_global tmp_thy rhs))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   439
    val _ = map check_rhs doms
40490
05be0c37db1d add bifiniteness check for domain_isomorphism command
huffman
parents: 40489
diff changeset
   440
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   441
    (* domain equations *)
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 44066
diff changeset
   442
    fun mk_dom_eqn (vs, tbind, _, rhs, _) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   443
      let fun arg v = TFree (v, the_sort v)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   444
      in (Type (Sign.full_name tmp_thy tbind, map arg vs), rhs) end
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   445
    val dom_eqns = map mk_dom_eqn doms
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   446
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   447
    (* check for valid type parameters *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   448
    val (tyvars, _, _, _, _) = hd doms
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 44066
diff changeset
   449
    val _ = map (fn (tvs, tname, _, _, _) =>
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   450
      let val full_tname = Sign.full_name tmp_thy tname
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   451
      in
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   452
        (case duplicates (op =) tvs of
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   453
          [] =>
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   454
            if eq_set (op =) (tyvars, tvs) then (full_tname, tvs)
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   455
            else error ("Mutually recursive domains must have same type parameters")
42381
309ec68442c6 added Binding.print convenience, which includes quote already;
wenzelm
parents: 42375
diff changeset
   456
        | dups => error ("Duplicate parameter(s) for domain " ^ Binding.print tname ^
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   457
            " : " ^ commas dups))
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   458
      end) doms
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   459
    val dbinds = map (fn (_, dbind, _, _, _) => dbind) doms
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   460
    val morphs = map (fn (_, _, _, _, morphs) => morphs) doms
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   461
40218
f7d4d023a899 make domain package work with non-cpo argument types
huffman
parents: 40216
diff changeset
   462
    (* determine deflation combinator arguments *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   463
    val lhsTs : typ list = map fst dom_eqns
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   464
    val defl_rec = Free ("t", mk_tupleT (map (K deflT) lhsTs))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   465
    val defl_recs = mk_projs lhsTs defl_rec
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   466
    val defl_recs' = map (apsnd mk_u_defl) defl_recs
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   467
    fun defl_body (_, _, _, rhsT, _) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   468
      defl_of_typ tmp_thy defl_recs defl_recs' rhsT
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   469
    val functional = Term.lambda defl_rec (mk_tuple (map defl_body doms))
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   470
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   471
    val tfrees = map fst (Term.add_tfrees functional [])
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   472
    val frees = map fst (Term.add_frees functional [])
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   473
    fun get_defl_flags (vs, _, _, _, _) =
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   474
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   475
        fun argT v = TFree (v, the_sort v)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   476
        fun mk_d v = "d" ^ Library.unprefix "'" v
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   477
        fun mk_p v = "p" ^ Library.unprefix "'" v
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   478
        val args = maps (fn v => [(mk_d v, mk_DEFL (argT v)), (mk_p v, mk_LIFTDEFL (argT v))]) vs
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   479
        val typeTs = map argT (filter (member (op =) tfrees) vs)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   480
        val defl_args = map snd (filter (member (op =) frees o fst) args)
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   481
      in
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   482
        (typeTs, defl_args)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   483
      end
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   484
    val defl_flagss = map get_defl_flags doms
40218
f7d4d023a899 make domain package work with non-cpo argument types
huffman
parents: 40216
diff changeset
   485
33784
7e434813752f change naming convention for deflation combinators
huffman
parents: 33782
diff changeset
   486
    (* declare deflation combinator constants *)
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   487
    fun declare_defl_const ((typeTs, defl_args), (_, tbind, _, _, _)) thy =
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   488
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   489
        val defl_bind = Binding.suffix_name "_defl" tbind
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   490
        val defl_type =
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41287
diff changeset
   491
          map Term.itselfT typeTs ---> map fastype_of defl_args -->> deflT
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   492
      in
42375
774df7c59508 report Name_Space.declare/define, relatively to context;
wenzelm
parents: 42361
diff changeset
   493
        Sign.declare_const_global ((defl_bind, defl_type), NoSyn) thy
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   494
      end
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   495
    val (defl_consts, thy) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   496
      fold_map declare_defl_const (defl_flagss ~~ doms) thy
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   497
33775
7a1518c42c56 cleaned up; factored out fixed-point definition code
huffman
parents: 33774
diff changeset
   498
    (* defining equations for type combinators *)
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   499
    fun mk_defl_term (defl_const, (typeTs, defl_args)) =
40487
1320a0747974 implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents: 40486
diff changeset
   500
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   501
        val type_args = map Logic.mk_type typeTs
40487
1320a0747974 implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents: 40486
diff changeset
   502
      in
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   503
        list_ccomb (list_comb (defl_const, type_args), defl_args)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   504
      end
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   505
    val defl_terms = map mk_defl_term (defl_consts ~~ defl_flagss)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   506
    val defl_tab = map fst dom_eqns ~~ defl_terms
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   507
    val defl_tab' = map fst dom_eqns ~~ map mk_u_defl defl_terms
33776
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   508
    fun mk_defl_spec (lhsT, rhsT) =
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   509
      mk_eqs (defl_of_typ tmp_thy defl_tab defl_tab' lhsT,
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   510
              defl_of_typ tmp_thy defl_tab defl_tab' rhsT)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   511
    val defl_specs = map mk_defl_spec dom_eqns
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   512
33784
7e434813752f change naming convention for deflation combinators
huffman
parents: 33782
diff changeset
   513
    (* register recursive definition of deflation combinators *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   514
    val defl_binds = map (Binding.suffix_name "_defl") dbinds
41324
1383653efec3 make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents: 41292
diff changeset
   515
    val ((defl_apply_thms, defl_unfold_thms, defl_cont_thm), thy) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   516
      add_fixdefs (defl_binds ~~ defl_specs) thy
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   517
33776
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   518
    (* define types using deflation combinators *)
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   519
    fun make_repdef ((vs, tbind, mx, _, _), defl) thy =
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   520
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   521
        val spec = (tbind, map (rpair dummyS) vs, mx)
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 44066
diff changeset
   522
        val ((_, _, _, {DEFL, ...}), thy) =
49759
ecf1d5d87e5e removed support for set constant definitions in HOLCF {cpo,pcpo,domain}def commands;
huffman
parents: 46961
diff changeset
   523
          Domaindef.add_domaindef spec defl NONE thy
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   524
        (* declare domain_defl_simps rules *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   525
        val thy = Context.theory_map (RepData.add_thm DEFL) thy
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   526
      in
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39974
diff changeset
   527
        (DEFL, thy)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   528
      end
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   529
    val (DEFL_thms, thy) = fold_map make_repdef (doms ~~ defl_terms) thy
33776
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   530
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39974
diff changeset
   531
    (* prove DEFL equations *)
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39974
diff changeset
   532
    fun mk_DEFL_eq_thm (lhsT, rhsT) =
33776
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   533
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   534
        val goal = mk_eqs (mk_DEFL lhsT, mk_DEFL rhsT)
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42151
diff changeset
   535
        val DEFL_simps = RepData.get (Proof_Context.init_global thy)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   536
        fun tac ctxt =
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   537
          rewrite_goals_tac ctxt (map mk_meta_eq DEFL_simps)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   538
          THEN TRY (resolve_tac defl_unfold_thms 1)
33776
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   539
      in
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   540
        Goal.prove_global thy [] [] goal (tac o #context)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   541
      end
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   542
    val DEFL_eq_thms = map mk_DEFL_eq_thm dom_eqns
33782
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   543
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39974
diff changeset
   544
    (* register DEFL equations *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   545
    val DEFL_eq_binds = map (Binding.prefix_name "DEFL_eq_") dbinds
33782
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   546
    val (_, thy) = thy |>
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37744
diff changeset
   547
      (Global_Theory.add_thms o map Thm.no_attributes)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   548
        (DEFL_eq_binds ~~ DEFL_eq_thms)
33776
5048b02c2bbb automate proofs of REP equations
huffman
parents: 33775
diff changeset
   549
33778
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   550
    (* define rep/abs functions *)
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 44066
diff changeset
   551
    fun mk_rep_abs ((tbind, _), (lhsT, rhsT)) thy =
33778
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   552
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   553
        val rep_bind = Binding.suffix_name "_rep" tbind
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   554
        val abs_bind = Binding.suffix_name "_abs" tbind
35477
b972233773dd add function define_const
huffman
parents: 35475
diff changeset
   555
        val ((rep_const, rep_def), thy) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   556
            define_const (rep_bind, coerce_const (lhsT, rhsT)) thy
35477
b972233773dd add function define_const
huffman
parents: 35475
diff changeset
   557
        val ((abs_const, abs_def), thy) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   558
            define_const (abs_bind, coerce_const (rhsT, lhsT)) thy
33778
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   559
      in
33785
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   560
        (((rep_const, abs_const), (rep_def, abs_def)), thy)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   561
      end
33785
2f2d9eb37084 automate definition of map functions; remove unused code
huffman
parents: 33784
diff changeset
   562
    val ((rep_abs_consts, rep_abs_defs), thy) = thy
35773
cae4f840d15d more consistent use of qualified bindings
huffman
parents: 35659
diff changeset
   563
      |> fold_map mk_rep_abs (dbinds ~~ morphs ~~ dom_eqns)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   564
      |>> ListPair.unzip
33782
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   565
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   566
    (* prove isomorphism and isodefl rules *)
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39974
diff changeset
   567
    fun mk_iso_thms ((tbind, DEFL_eq), (rep_def, abs_def)) thy =
33782
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   568
      let
35773
cae4f840d15d more consistent use of qualified bindings
huffman
parents: 35659
diff changeset
   569
        fun make thm =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   570
            Drule.zero_var_indexes (thm OF [DEFL_eq, abs_def, rep_def])
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   571
        val rep_iso_thm = make @{thm domain_rep_iso}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   572
        val abs_iso_thm = make @{thm domain_abs_iso}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   573
        val isodefl_thm = make @{thm isodefl_abs_rep}
35773
cae4f840d15d more consistent use of qualified bindings
huffman
parents: 35659
diff changeset
   574
        val thy = thy
cae4f840d15d more consistent use of qualified bindings
huffman
parents: 35659
diff changeset
   575
          |> snd o add_qualified_thm "rep_iso" (tbind, rep_iso_thm)
cae4f840d15d more consistent use of qualified bindings
huffman
parents: 35659
diff changeset
   576
          |> snd o add_qualified_thm "abs_iso" (tbind, abs_iso_thm)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   577
          |> snd o add_qualified_thm "isodefl_abs_rep" (tbind, isodefl_thm)
33782
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   578
      in
cdb3ca1a765d prove isomorphism and isodefl rules
huffman
parents: 33778
diff changeset
   579
        (((rep_iso_thm, abs_iso_thm), isodefl_thm), thy)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   580
      end
35489
dd02201d95b6 add function define_take_functions
huffman
parents: 35480
diff changeset
   581
    val ((iso_thms, isodefl_abs_rep_thms), thy) =
35480
7a1f285cad25 domain_isomorphism package proves deflation rules for map functions
huffman
parents: 35479
diff changeset
   582
      thy
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39974
diff changeset
   583
      |> fold_map mk_iso_thms (dbinds ~~ DEFL_eq_thms ~~ rep_abs_defs)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   584
      |>> ListPair.unzip
33778
9121ea165a40 automate definition of rep/abs functions
huffman
parents: 33777
diff changeset
   585
35467
561d8e98d9d3 domain_isomorphism function returns iso_info record
huffman
parents: 35021
diff changeset
   586
    (* collect info about rep/abs *)
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents: 35498
diff changeset
   587
    val iso_infos : Domain_Take_Proofs.iso_info list =
35467
561d8e98d9d3 domain_isomorphism function returns iso_info record
huffman
parents: 35021
diff changeset
   588
      let
561d8e98d9d3 domain_isomorphism function returns iso_info record
huffman
parents: 35021
diff changeset
   589
        fun mk_info (((lhsT, rhsT), (repC, absC)), (rep_iso, abs_iso)) =
561d8e98d9d3 domain_isomorphism function returns iso_info record
huffman
parents: 35021
diff changeset
   590
          {
561d8e98d9d3 domain_isomorphism function returns iso_info record
huffman
parents: 35021
diff changeset
   591
            repT = rhsT,
561d8e98d9d3 domain_isomorphism function returns iso_info record
huffman
parents: 35021
diff changeset
   592
            absT = lhsT,
561d8e98d9d3 domain_isomorphism function returns iso_info record
huffman
parents: 35021
diff changeset
   593
            rep_const = repC,
561d8e98d9d3 domain_isomorphism function returns iso_info record
huffman
parents: 35021
diff changeset
   594
            abs_const = absC,
561d8e98d9d3 domain_isomorphism function returns iso_info record
huffman
parents: 35021
diff changeset
   595
            rep_inverse = rep_iso,
561d8e98d9d3 domain_isomorphism function returns iso_info record
huffman
parents: 35021
diff changeset
   596
            abs_inverse = abs_iso
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   597
          }
35467
561d8e98d9d3 domain_isomorphism function returns iso_info record
huffman
parents: 35021
diff changeset
   598
      in
561d8e98d9d3 domain_isomorphism function returns iso_info record
huffman
parents: 35021
diff changeset
   599
        map mk_info (dom_eqns ~~ rep_abs_consts ~~ iso_thms)
561d8e98d9d3 domain_isomorphism function returns iso_info record
huffman
parents: 35021
diff changeset
   600
      end
561d8e98d9d3 domain_isomorphism function returns iso_info record
huffman
parents: 35021
diff changeset
   601
35791
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   602
    (* definitions and proofs related to map functions *)
dc175fe29326 separate map-related code into new function define_map_functions
huffman
parents: 35773
diff changeset
   603
    val (map_info, thy) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   604
        define_map_functions (dbinds ~~ iso_infos) thy
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 44066
diff changeset
   605
    val { map_consts, map_apply_thms, map_cont_thm, ...} = map_info
33788
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   606
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   607
    (* prove isodefl rules for map functions *)
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   608
    val isodefl_thm =
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   609
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   610
        fun unprime a = Library.unprefix "'" a
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   611
        fun mk_d T = Free ("d" ^ unprime (fst (dest_TFree T)), deflT)
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41287
diff changeset
   612
        fun mk_p T = Free ("p" ^ unprime (fst (dest_TFree T)), udeflT)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   613
        fun mk_f T = Free ("f" ^ unprime (fst (dest_TFree T)), T ->> T)
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   614
        fun mk_assm t =
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   615
          case try dest_LIFTDEFL t of
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41287
diff changeset
   616
            SOME T => mk_trp (isodefl'_const T $ mk_f T $ mk_p T)
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   617
          | NONE =>
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   618
            let val T = dest_DEFL t
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   619
            in mk_trp (isodefl_const T $ mk_f T $ mk_d T) end
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 44066
diff changeset
   620
        fun mk_goal (map_const, (T, _)) =
33788
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   621
          let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   622
            val (_, Ts) = dest_Type T
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   623
            val map_term = list_ccomb (map_const, map mk_f (filter (is_cpo thy) Ts))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   624
            val defl_term = defl_of_typ thy (Ts ~~ map mk_d Ts) (Ts ~~ map mk_p Ts) T
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   625
          in isodefl_const T $ map_term $ defl_term end
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   626
        val assms = (map mk_assm o snd o hd) defl_flagss
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   627
        val goals = map mk_goal (map_consts ~~ dom_eqns)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   628
        val goal = mk_trp (foldr1 HOLogic.mk_conj goals)
33788
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   629
        val adm_rules =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   630
          @{thms adm_conj adm_isodefl cont2cont_fst cont2cont_snd cont_id}
33788
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   631
        val bottom_rules =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   632
          @{thms fst_strict snd_strict isodefl_bottom simp_thms}
41324
1383653efec3 make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents: 41292
diff changeset
   633
        val tuple_rules =
1383653efec3 make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents: 41292
diff changeset
   634
          @{thms split_def fst_conv snd_conv}
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   635
        val map_ID_thms = Domain_Take_Proofs.get_map_ID_thms thy
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   636
        val map_ID_simps = map (fn th => th RS sym) map_ID_thms
33788
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   637
        val isodefl_rules =
40491
6de5839e2fb3 add 'predomain' class: unpointed version of bifinite
huffman
parents: 40490
diff changeset
   638
          @{thms conjI isodefl_ID_DEFL isodefl_LIFTDEFL}
33791
fef59343b4b3 use theory data for REP_simps and isodefl_rules
huffman
parents: 33790
diff changeset
   639
          @ isodefl_abs_rep_thms
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42151
diff changeset
   640
          @ IsodeflData.get (Proof_Context.init_global thy)
33845
91f3fc0364cf made SML/NJ happy;
wenzelm
parents: 33809
diff changeset
   641
      in
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   642
        Goal.prove_global thy [] assms goal (fn {prems, context = ctxt} =>
33845
91f3fc0364cf made SML/NJ happy;
wenzelm
parents: 33809
diff changeset
   643
         EVERY
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52788
diff changeset
   644
          [rewrite_goals_tac ctxt (defl_apply_thms @ map_apply_thms),
41324
1383653efec3 make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents: 41292
diff changeset
   645
           rtac (@{thm cont_parallel_fix_ind}
1383653efec3 make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents: 41292
diff changeset
   646
             OF [defl_cont_thm, map_cont_thm]) 1,
33788
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   647
           REPEAT (resolve_tac adm_rules 1),
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   648
           simp_tac (put_simpset HOL_basic_ss ctxt addsimps bottom_rules) 1,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   649
           simp_tac (put_simpset HOL_basic_ss ctxt addsimps tuple_rules) 1,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   650
           simp_tac (put_simpset HOL_basic_ss ctxt addsimps map_ID_simps) 1,
33788
481bc899febf automate isodefl proof
huffman
parents: 33785
diff changeset
   651
           REPEAT (etac @{thm conjE} 1),
33845
91f3fc0364cf made SML/NJ happy;
wenzelm
parents: 33809
diff changeset
   652
           REPEAT (resolve_tac (isodefl_rules @ prems) 1 ORELSE atac 1)])
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   653
      end
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   654
    val isodefl_binds = map (Binding.prefix_name "isodefl_") dbinds
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 44066
diff changeset
   655
    fun conjuncts [] _ = []
33789
c3fbdff7aed0 separate conjuncts of isodefl theorem
huffman
parents: 33788
diff changeset
   656
      | conjuncts (n::[]) thm = [(n, thm)]
c3fbdff7aed0 separate conjuncts of isodefl theorem
huffman
parents: 33788
diff changeset
   657
      | conjuncts (n::ns) thm = let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   658
          val thmL = thm RS @{thm conjunct1}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   659
          val thmR = thm RS @{thm conjunct2}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   660
        in (n, thmL):: conjuncts ns thmR end
33789
c3fbdff7aed0 separate conjuncts of isodefl theorem
huffman
parents: 33788
diff changeset
   661
    val (isodefl_thms, thy) = thy |>
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37744
diff changeset
   662
      (Global_Theory.add_thms o map (Thm.no_attributes o apsnd Drule.zero_var_indexes))
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   663
        (conjuncts isodefl_binds isodefl_thm)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   664
    val thy = fold (Context.theory_map o IsodeflData.add_thm) isodefl_thms thy
33793
5beafabffa07 automate proofs of map_ID theorems
huffman
parents: 33792
diff changeset
   665
5beafabffa07 automate proofs of map_ID theorems
huffman
parents: 33792
diff changeset
   666
    (* prove map_ID theorems *)
5beafabffa07 automate proofs of map_ID theorems
huffman
parents: 33792
diff changeset
   667
    fun prove_map_ID_thm
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39974
diff changeset
   668
        (((map_const, (lhsT, _)), DEFL_thm), isodefl_thm) =
33793
5beafabffa07 automate proofs of map_ID theorems
huffman
parents: 33792
diff changeset
   669
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   670
        val Ts = snd (dest_Type lhsT)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   671
        fun is_cpo T = Sign.of_sort thy (T, @{sort cpo})
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   672
        val lhs = list_ccomb (map_const, map mk_ID (filter is_cpo Ts))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   673
        val goal = mk_eqs (lhs, mk_ID lhsT)
33793
5beafabffa07 automate proofs of map_ID theorems
huffman
parents: 33792
diff changeset
   674
        val tac = EVERY
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39974
diff changeset
   675
          [rtac @{thm isodefl_DEFL_imp_ID} 1,
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39974
diff changeset
   676
           stac DEFL_thm 1,
33793
5beafabffa07 automate proofs of map_ID theorems
huffman
parents: 33792
diff changeset
   677
           rtac isodefl_thm 1,
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   678
           REPEAT (resolve_tac @{thms isodefl_ID_DEFL isodefl_LIFTDEFL} 1)]
33793
5beafabffa07 automate proofs of map_ID theorems
huffman
parents: 33792
diff changeset
   679
      in
5beafabffa07 automate proofs of map_ID theorems
huffman
parents: 33792
diff changeset
   680
        Goal.prove_global thy [] [] goal (K tac)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   681
      end
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   682
    val map_ID_binds = map (Binding.suffix_name "_map_ID") dbinds
33793
5beafabffa07 automate proofs of map_ID theorems
huffman
parents: 33792
diff changeset
   683
    val map_ID_thms =
5beafabffa07 automate proofs of map_ID theorems
huffman
parents: 33792
diff changeset
   684
      map prove_map_ID_thm
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   685
        (map_consts ~~ dom_eqns ~~ DEFL_thms ~~ isodefl_thms)
33793
5beafabffa07 automate proofs of map_ID theorems
huffman
parents: 33792
diff changeset
   686
    val (_, thy) = thy |>
40489
0f37a553bc1a implement map_of_typ using Pattern.rewrite_term
huffman
parents: 40487
diff changeset
   687
      (Global_Theory.add_thms o map (rpair [Domain_Take_Proofs.map_ID_add]))
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   688
        (map_ID_binds ~~ map_ID_thms)
33802
48ce3a1063f2 domain_isomorphism package defines copy functions
huffman
parents: 33801
diff changeset
   689
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35490
diff changeset
   690
    (* definitions and proofs related to take functions *)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35490
diff changeset
   691
    val (take_info, thy) =
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents: 35498
diff changeset
   692
        Domain_Take_Proofs.define_take_functions
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   693
          (dbinds ~~ iso_infos) thy
40015
2fda96749081 add take_strict_thms field to take_info type
huffman
parents: 39989
diff changeset
   694
    val { take_consts, chain_take_thms, take_0_thms, take_Suc_thms, ...} =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   695
        take_info
33807
ce8d2e8bca21 domain_isomorphism package defines combined copy function
huffman
parents: 33802
diff changeset
   696
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35490
diff changeset
   697
    (* least-upper-bound lemma for take functions *)
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35490
diff changeset
   698
    val lub_take_lemma =
33809
033831fd9ef3 store map_ID thms in theory data; automate proofs of reach lemmas
huffman
parents: 33807
diff changeset
   699
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   700
        val lhs = mk_tuple (map mk_lub take_consts)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   701
        fun is_cpo T = Sign.of_sort thy (T, @{sort cpo})
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 44066
diff changeset
   702
        fun mk_map_ID (map_const, (lhsT, _)) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   703
          list_ccomb (map_const, map mk_ID (filter is_cpo (snd (dest_Type lhsT))))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   704
        val rhs = mk_tuple (map mk_map_ID (map_consts ~~ dom_eqns))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   705
        val goal = mk_trp (mk_eq (lhs, rhs))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   706
        val map_ID_thms = Domain_Take_Proofs.get_map_ID_thms thy
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35490
diff changeset
   707
        val start_rules =
40771
1c6f7d4b110e renamed several HOLCF theorems (listed in NEWS)
huffman
parents: 40737
diff changeset
   708
            @{thms lub_Pair [symmetric] ch2ch_Pair} @ chain_take_thms
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35490
diff changeset
   709
            @ @{thms pair_collapse split_def}
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   710
            @ map_apply_thms @ map_ID_thms
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35490
diff changeset
   711
        val rules0 =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   712
            @{thms iterate_0 Pair_strict} @ take_0_thms
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35490
diff changeset
   713
        val rules1 =
44066
d74182c93f04 rename Pair_fst_snd_eq to prod_eq_iff (keeping old name too)
huffman
parents: 42381
diff changeset
   714
            @{thms iterate_Suc prod_eq_iff fst_conv snd_conv}
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   715
            @ take_Suc_thms
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   716
        fun tac ctxt =
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35490
diff changeset
   717
            EVERY
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   718
            [simp_tac (put_simpset HOL_basic_ss ctxt addsimps start_rules) 1,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   719
             simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms fix_def2}) 1,
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35490
diff changeset
   720
             rtac @{thm lub_eq} 1,
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35490
diff changeset
   721
             rtac @{thm nat.induct} 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   722
             simp_tac (put_simpset HOL_basic_ss ctxt addsimps rules0) 1,
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   723
             asm_full_simp_tac (put_simpset beta_ss ctxt addsimps rules1) 1]
33809
033831fd9ef3 store map_ID thms in theory data; automate proofs of reach lemmas
huffman
parents: 33807
diff changeset
   724
      in
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51685
diff changeset
   725
        Goal.prove_global thy [] [] goal (tac o #context)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   726
      end
33809
033831fd9ef3 store map_ID thms in theory data; automate proofs of reach lemmas
huffman
parents: 33807
diff changeset
   727
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35490
diff changeset
   728
    (* prove lub of take equals ID *)
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 44066
diff changeset
   729
    fun prove_lub_take (((dbind, take_const), map_ID_thm), (lhsT, _)) thy =
33809
033831fd9ef3 store map_ID thms in theory data; automate proofs of reach lemmas
huffman
parents: 33807
diff changeset
   730
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   731
        val n = Free ("n", natT)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   732
        val goal = mk_eqs (mk_lub (lambda n (take_const $ n)), mk_ID lhsT)
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35490
diff changeset
   733
        val tac =
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35490
diff changeset
   734
            EVERY
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35490
diff changeset
   735
            [rtac @{thm trans} 1, rtac map_ID_thm 2,
46708
b138dee7bed3 prefer cut_tac, where it is clear that the special variants cut_rules_tac or cut_facts_tac are not required;
wenzelm
parents: 45654
diff changeset
   736
             cut_tac lub_take_lemma 1,
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   737
             REPEAT (etac @{thm Pair_inject} 1), atac 1]
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   738
        val lub_take_thm = Goal.prove_global thy [] [] goal (K tac)
33809
033831fd9ef3 store map_ID thms in theory data; automate proofs of reach lemmas
huffman
parents: 33807
diff changeset
   739
      in
35773
cae4f840d15d more consistent use of qualified bindings
huffman
parents: 35659
diff changeset
   740
        add_qualified_thm "lub_take" (dbind, lub_take_thm) thy
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   741
      end
35494
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35490
diff changeset
   742
    val (lub_take_thms, thy) =
45c9a8278faf domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents: 35490
diff changeset
   743
        fold_map prove_lub_take
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   744
          (dbinds ~~ take_consts ~~ map_ID_thms ~~ dom_eqns) thy
33809
033831fd9ef3 store map_ID thms in theory data; automate proofs of reach lemmas
huffman
parents: 33807
diff changeset
   745
35654
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35640
diff changeset
   746
    (* prove additional take theorems *)
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35640
diff changeset
   747
    val (take_info2, thy) =
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35640
diff changeset
   748
        Domain_Take_Proofs.add_lub_take_theorems
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   749
          (dbinds ~~ iso_infos) take_info lub_take_thms thy
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   750
  in
35659
a78bc1930a7a include take_info within take_induct_info type
huffman
parents: 35658
diff changeset
   751
    ((iso_infos, take_info2), thy)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   752
  end
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   753
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   754
val domain_isomorphism = gen_domain_isomorphism cert_typ
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   755
val domain_isomorphism_cmd = snd oo gen_domain_isomorphism read_typ
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   756
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   757
(******************************************************************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   758
(******************************** outer syntax ********************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   759
(******************************************************************************)
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   760
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   761
local
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   762
34149
a0efb4754cb7 add 'morphisms' option to domain_isomorphism command
huffman
parents: 33971
diff changeset
   763
val parse_domain_iso :
a0efb4754cb7 add 'morphisms' option to domain_isomorphism command
huffman
parents: 33971
diff changeset
   764
    (string list * binding * mixfix * string * (binding * binding) option)
a0efb4754cb7 add 'morphisms' option to domain_isomorphism command
huffman
parents: 33971
diff changeset
   765
      parser =
46949
94aa7b81bcf6 prefer formally checked @{keyword} parser;
wenzelm
parents: 46909
diff changeset
   766
  (Parse.type_args -- Parse.binding -- Parse.opt_mixfix -- (@{keyword "="} |-- Parse.typ) --
94aa7b81bcf6 prefer formally checked @{keyword} parser;
wenzelm
parents: 46909
diff changeset
   767
    Scan.option (@{keyword "morphisms"} |-- Parse.!!! (Parse.binding -- Parse.binding)))
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   768
    >> (fn ((((vs, t), mx), rhs), morphs) => (vs, t, mx, rhs, morphs))
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   769
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   770
val parse_domain_isos = Parse.and_list1 parse_domain_iso
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   771
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   772
in
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   773
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   774
val _ =
46961
5c6955f487e5 outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents: 46949
diff changeset
   775
  Outer_Syntax.command @{command_spec "domain_isomorphism"} "define domain isomorphisms (HOLCF)"
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   776
    (parse_domain_isos >> (Toplevel.theory o domain_isomorphism_cmd))
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   777
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   778
end
33774
e11e05b32548 automate solution of domain equations
huffman
parents:
diff changeset
   779
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   780
end