src/HOL/HOLCF/Tools/Domain/domain_constructors.ML
author blanchet
Mon, 01 Sep 2014 16:17:46 +0200
changeset 58112 8081087096ad
parent 54895 515630483010
child 59058 a78612c67ec0
permissions -rw-r--r--
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 42057
diff changeset
     1
(*  Title:      HOL/HOLCF/Tools/Domain/domain_constructors.ML
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
     2
    Author:     Brian Huffman
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
     3
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
     4
Defines constructor functions for a given domain isomorphism
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
     5
and proves related theorems.
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
     6
*)
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
     7
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
     8
signature DOMAIN_CONSTRUCTORS =
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
     9
sig
40014
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    10
  type constr_info =
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    11
    {
40017
575d3aa1f3c5 include iso_info as part of constr_info type
huffman
parents: 40016
diff changeset
    12
      iso_info : Domain_Take_Proofs.iso_info,
40019
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40017
diff changeset
    13
      con_specs : (term * (bool * typ) list) list,
40014
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    14
      con_betas : thm list,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    15
      nchotomy : thm,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    16
      exhaust : thm,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    17
      compacts : thm list,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    18
      con_rews : thm list,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    19
      inverts : thm list,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    20
      injects : thm list,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    21
      dist_les : thm list,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    22
      dist_eqs : thm list,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    23
      cases : thm list,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    24
      sel_rews : thm list,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    25
      dis_rews : thm list,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    26
      match_rews : thm list
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    27
    }
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
    28
  val add_domain_constructors :
35777
bcc77916b7b9 pass binding as argument to add_domain_constructors; proper binding for case combinator
huffman
parents: 35561
diff changeset
    29
      binding
35481
7bb9157507a9 add_domain_constructors takes iso_info record as argument
huffman
parents: 35476
diff changeset
    30
      -> (binding * (bool * binding option * typ) list * mixfix) list
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents: 35496
diff changeset
    31
      -> Domain_Take_Proofs.iso_info
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
    32
      -> theory
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    33
      -> constr_info * theory
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    34
end
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
    35
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
    36
41296
6aaf80ea9715 switch to transparent ascription, to avoid warning messages
huffman
parents: 40853
diff changeset
    37
structure Domain_Constructors : DOMAIN_CONSTRUCTORS =
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
    38
struct
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
    39
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    40
open HOLCF_Library
35561
b56d5b1b1a55 add infix declarations
huffman
parents: 35526
diff changeset
    41
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    42
infixr 6 ->>
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    43
infix -->>
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    44
infix 9 `
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
    45
40014
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    46
type constr_info =
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    47
  {
40017
575d3aa1f3c5 include iso_info as part of constr_info type
huffman
parents: 40016
diff changeset
    48
    iso_info : Domain_Take_Proofs.iso_info,
40019
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40017
diff changeset
    49
    con_specs : (term * (bool * typ) list) list,
40014
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    50
    con_betas : thm list,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    51
    nchotomy : thm,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    52
    exhaust : thm,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    53
    compacts : thm list,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    54
    con_rews : thm list,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    55
    inverts : thm list,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    56
    injects : thm list,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    57
    dist_les : thm list,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    58
    dist_eqs : thm list,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    59
    cases : thm list,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    60
    sel_rews : thm list,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    61
    dis_rews : thm list,
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    62
    match_rews : thm list
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    63
  }
7469b323e260 add record type synonym 'constr_info'
huffman
parents: 39557
diff changeset
    64
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
    65
(************************** miscellaneous functions ***************************)
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
    66
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48563
diff changeset
    67
val simple_ss =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48563
diff changeset
    68
  simpset_of (put_simpset HOL_basic_ss @{context} addsimps @{thms simp_thms})
35456
5356534f880e move proofs of injects and inverts to domain_constructors.ML
huffman
parents: 35455
diff changeset
    69
37078
a1656804fcad domain package internal proofs use fixed set of continuity rules, rather than taking cont2cont rules from context
huffman
parents: 36998
diff changeset
    70
val beta_rules =
40326
73d45866dbda renamed lemma cont2cont_Rep_CFun to cont2cont_APP
huffman
parents: 40321
diff changeset
    71
  @{thms beta_cfun cont_id cont_const cont2cont_APP cont2cont_LAM'} @
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    72
  @{thms cont2cont_fst cont2cont_snd cont2cont_Pair}
37078
a1656804fcad domain package internal proofs use fixed set of continuity rules, rather than taking cont2cont rules from context
huffman
parents: 36998
diff changeset
    73
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48563
diff changeset
    74
val beta_ss =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 48563
diff changeset
    75
  simpset_of (put_simpset HOL_basic_ss @{context} addsimps (@{thms simp_thms} @ beta_rules))
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
    76
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
    77
fun define_consts
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
    78
    (specs : (binding * term * mixfix) list)
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
    79
    (thy : theory)
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
    80
    : (term list * thm list) * theory =
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
    81
  let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    82
    fun mk_decl (b, t, mx) = (b, fastype_of t, mx)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    83
    val decls = map mk_decl specs
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    84
    val thy = Cont_Consts.add_consts decls thy
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
    85
    fun mk_const (b, T, _) = Const (Sign.full_name thy b, T)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    86
    val consts = map mk_const decls
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
    87
    fun mk_def c (b, t, _) =
46909
3c73a121a387 more explicit indication of def names;
wenzelm
parents: 46490
diff changeset
    88
      (Thm.def_binding b, Logic.mk_equals (c, t))
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    89
    val defs = map2 mk_def consts specs
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
    90
    val (def_thms, thy) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    91
      Global_Theory.add_defs false (map Thm.no_attributes defs) thy
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
    92
  in
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
    93
    ((consts, def_thms), thy)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
    94
  end
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
    95
35449
1d6657074fcb replace prove_thm function
huffman
parents: 35448
diff changeset
    96
fun prove
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
    97
    (thy : theory)
35449
1d6657074fcb replace prove_thm function
huffman
parents: 35448
diff changeset
    98
    (defs : thm list)
35447
82af95d998e0 add some convenience functions
huffman
parents: 35446
diff changeset
    99
    (goal : term)
35449
1d6657074fcb replace prove_thm function
huffman
parents: 35448
diff changeset
   100
    (tacs : {prems: thm list, context: Proof.context} -> tactic list)
35447
82af95d998e0 add some convenience functions
huffman
parents: 35446
diff changeset
   101
    : thm =
35449
1d6657074fcb replace prove_thm function
huffman
parents: 35448
diff changeset
   102
  let
1d6657074fcb replace prove_thm function
huffman
parents: 35448
diff changeset
   103
    fun tac {prems, context} =
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   104
      rewrite_goals_tac context defs THEN
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   105
      EVERY (tacs {prems = map (rewrite_rule context defs) prems, context = context})
35449
1d6657074fcb replace prove_thm function
huffman
parents: 35448
diff changeset
   106
  in
1d6657074fcb replace prove_thm function
huffman
parents: 35448
diff changeset
   107
    Goal.prove_global thy [] [] goal tac
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   108
  end
35445
593c184977a4 add function beta_of_def
huffman
parents: 35444
diff changeset
   109
35483
614b42e719ee add functions get_vars, get_vars_avoiding
huffman
parents: 35482
diff changeset
   110
fun get_vars_avoiding
614b42e719ee add functions get_vars, get_vars_avoiding
huffman
parents: 35482
diff changeset
   111
    (taken : string list)
614b42e719ee add functions get_vars, get_vars_avoiding
huffman
parents: 35482
diff changeset
   112
    (args : (bool * typ) list)
614b42e719ee add functions get_vars, get_vars_avoiding
huffman
parents: 35482
diff changeset
   113
    : (term list * term list) =
614b42e719ee add functions get_vars, get_vars_avoiding
huffman
parents: 35482
diff changeset
   114
  let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   115
    val Ts = map snd args
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 54895
diff changeset
   116
    val ns = Name.variant_list taken (Old_Datatype_Prop.make_tnames Ts)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   117
    val vs = map Free (ns ~~ Ts)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   118
    val nonlazy = map snd (filter_out (fst o fst) (args ~~ vs))
35483
614b42e719ee add functions get_vars, get_vars_avoiding
huffman
parents: 35482
diff changeset
   119
  in
614b42e719ee add functions get_vars, get_vars_avoiding
huffman
parents: 35482
diff changeset
   120
    (vs, nonlazy)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   121
  end
35483
614b42e719ee add functions get_vars, get_vars_avoiding
huffman
parents: 35482
diff changeset
   122
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   123
fun get_vars args = get_vars_avoiding [] args
35483
614b42e719ee add functions get_vars, get_vars_avoiding
huffman
parents: 35482
diff changeset
   124
35445
593c184977a4 add function beta_of_def
huffman
parents: 35444
diff changeset
   125
(************** generating beta reduction rules from definitions **************)
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
   126
35445
593c184977a4 add function beta_of_def
huffman
parents: 35444
diff changeset
   127
local
593c184977a4 add function beta_of_def
huffman
parents: 35444
diff changeset
   128
  fun arglist (Const _ $ Abs (s, T, t)) =
593c184977a4 add function beta_of_def
huffman
parents: 35444
diff changeset
   129
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   130
        val arg = Free (s, T)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   131
        val (args, body) = arglist (subst_bound (arg, t))
35445
593c184977a4 add function beta_of_def
huffman
parents: 35444
diff changeset
   132
      in (arg :: args, body) end
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   133
    | arglist t = ([], t)
35445
593c184977a4 add function beta_of_def
huffman
parents: 35444
diff changeset
   134
in
593c184977a4 add function beta_of_def
huffman
parents: 35444
diff changeset
   135
  fun beta_of_def thy def_thm =
593c184977a4 add function beta_of_def
huffman
parents: 35444
diff changeset
   136
      let
48563
04e129931181 unvarify thm statement stemming from old-style definition, to avoid schematic type variables in subsequent goal;
wenzelm
parents: 46909
diff changeset
   137
        val (con, lam) =
04e129931181 unvarify thm statement stemming from old-style definition, to avoid schematic type variables in subsequent goal;
wenzelm
parents: 46909
diff changeset
   138
          Logic.dest_equals (Logic.unvarify_global (concl_of def_thm))
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   139
        val (args, rhs) = arglist lam
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   140
        val lhs = list_ccomb (con, args)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   141
        val goal = mk_equals (lhs, rhs)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   142
        val cs = ContProc.cont_thms lam
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   143
        val betas = map (fn c => mk_meta_eq (c RS @{thm beta_cfun})) cs
35445
593c184977a4 add function beta_of_def
huffman
parents: 35444
diff changeset
   144
      in
35449
1d6657074fcb replace prove_thm function
huffman
parents: 35448
diff changeset
   145
        prove thy (def_thm::betas) goal (K [rtac reflexive_thm 1])
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   146
      end
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   147
end
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
   148
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   149
(******************************************************************************)
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   150
(************* definitions and theorems for constructor functions *************)
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   151
(******************************************************************************)
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   152
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   153
fun add_constructors
35454
cf6ba350b7ca remove unnecessary stuff from argument to add_constructors function
huffman
parents: 35453
diff changeset
   154
    (spec : (binding * (bool * typ) list * mixfix) list)
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   155
    (abs_const : term)
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   156
    (iso_locale : thm)
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   157
    (thy : theory)
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   158
    =
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   159
  let
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   160
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   161
    (* get theorems about rep and abs *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   162
    val abs_strict = iso_locale RS @{thm iso.abs_strict}
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   163
35457
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   164
    (* get types of type isomorphism *)
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   165
    val (_, lhsT) = dest_cfunT (fastype_of abs_const)
35457
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   166
35455
490a6e6c7379 reuse vars_of function
huffman
parents: 35454
diff changeset
   167
    fun vars_of args =
490a6e6c7379 reuse vars_of function
huffman
parents: 35454
diff changeset
   168
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   169
        val Ts = map snd args
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 54895
diff changeset
   170
        val ns = Old_Datatype_Prop.make_tnames Ts
35455
490a6e6c7379 reuse vars_of function
huffman
parents: 35454
diff changeset
   171
      in
490a6e6c7379 reuse vars_of function
huffman
parents: 35454
diff changeset
   172
        map Free (ns ~~ Ts)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   173
      end
35455
490a6e6c7379 reuse vars_of function
huffman
parents: 35454
diff changeset
   174
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   175
    (* define constructor functions *)
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   176
    val ((con_consts, con_defs), thy) =
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   177
      let
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   178
        fun one_arg (lazy, _) var = if lazy then mk_up var else var
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   179
        fun one_con (_,args,_) = mk_stuple (map2 one_arg args (vars_of args))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   180
        fun mk_abs t = abs_const ` t
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   181
        val rhss = map mk_abs (mk_sinjects (map one_con spec))
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   182
        fun mk_def (bind, args, mx) rhs =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   183
          (bind, big_lambdas (vars_of args) rhs, mx)
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   184
      in
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   185
        define_consts (map2 mk_def spec rhss) thy
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   186
      end
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   187
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   188
    (* prove beta reduction rules for constructors *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   189
    val con_betas = map (beta_of_def thy) con_defs
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   190
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   191
    (* replace bindings with terms in constructor spec *)
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   192
    val spec' : (term * (bool * typ) list) list =
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   193
      let fun one_con con (_, args, _) = (con, args)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   194
      in map2 one_con con_consts spec end
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   195
35457
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   196
    (* prove exhaustiveness of constructors *)
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   197
    local
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   198
      fun arg2typ n (true,  _) = (n+1, mk_upT (TVar (("'a", n), @{sort cpo})))
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   199
        | arg2typ n (false, _) = (n+1, TVar (("'a", n), @{sort pcpo}))
35457
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   200
      fun args2typ n [] = (n, oneT)
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   201
        | args2typ n [arg] = arg2typ n arg
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   202
        | args2typ n (arg::args) =
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   203
          let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   204
            val (n1, t1) = arg2typ n arg
35457
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   205
            val (n2, t2) = args2typ n1 args
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   206
          in (n2, mk_sprodT (t1, t2)) end
35457
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   207
      fun cons2typ n [] = (n, oneT)
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   208
        | cons2typ n [con] = args2typ n (snd con)
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   209
        | cons2typ n (con::cons) =
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   210
          let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   211
            val (n1, t1) = args2typ n (snd con)
35457
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   212
            val (n2, t2) = cons2typ n1 cons
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   213
          in (n2, mk_ssumT (t1, t2)) end
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   214
      val ct = ctyp_of thy (snd (cons2typ 1 spec'))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   215
      val thm1 = instantiate' [SOME ct] [] @{thm exh_start}
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   216
      val thm2 = rewrite_rule (Proof_Context.init_global thy)
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   217
        (map mk_meta_eq @{thms ex_bottom_iffs}) thm1
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   218
      val thm3 = rewrite_rule (Proof_Context.init_global thy)
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   219
        [mk_meta_eq @{thm conj_assoc}] thm2
35457
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   220
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   221
      val y = Free ("y", lhsT)
35457
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   222
      fun one_con (con, args) =
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   223
        let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   224
          val (vs, nonlazy) = get_vars_avoiding ["y"] args
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   225
          val eqn = mk_eq (y, list_ccomb (con, vs))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   226
          val conj = foldr1 mk_conj (eqn :: map mk_defined nonlazy)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   227
        in Library.foldr mk_ex (vs, conj) end
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   228
      val goal = mk_trp (foldr1 mk_disj (mk_undef y :: map one_con spec'))
41429
cf5f025bc3c7 renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
huffman
parents: 41296
diff changeset
   229
      (* first rules replace "y = bottom \/ P" with "rep$y = bottom \/ P" *)
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   230
      fun tacs ctxt = [
35457
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   231
          rtac (iso_locale RS @{thm iso.casedist_rule}) 1,
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   232
          rewrite_goals_tac ctxt [mk_meta_eq (iso_locale RS @{thm iso.iso_swap})],
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   233
          rtac thm3 1]
35457
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   234
    in
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   235
      val nchotomy = prove thy con_betas goal (tacs o #context)
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   236
      val exhaust =
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   237
          (nchotomy RS @{thm exh_casedist0})
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   238
          |> rewrite_rule (Proof_Context.init_global thy) @{thms exh_casedists}
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   239
          |> Drule.zero_var_indexes
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   240
    end
35457
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   241
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   242
    (* prove compactness rules for constructors *)
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   243
    val compacts =
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   244
      let
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   245
        val rules = @{thms compact_sinl compact_sinr compact_spair
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   246
                           compact_up compact_ONE}
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   247
        val tacs =
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   248
          [rtac (iso_locale RS @{thm iso.compact_abs}) 1,
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   249
           REPEAT (resolve_tac rules 1 ORELSE atac 1)]
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   250
        fun con_compact (con, args) =
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   251
          let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   252
            val vs = vars_of args
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   253
            val con_app = list_ccomb (con, vs)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   254
            val concl = mk_trp (mk_compact con_app)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   255
            val assms = map (mk_trp o mk_compact) vs
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   256
            val goal = Logic.list_implies (assms, concl)
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   257
          in
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   258
            prove thy con_betas goal (K tacs)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   259
          end
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   260
      in
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   261
        map con_compact spec'
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   262
      end
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   263
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   264
    (* prove strictness rules for constructors *)
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   265
    local
46490
e4863ab5e09b more basic simplification, eliminated slightly odd tactic style from 1995 (cf. ea0668a1c0ba);
wenzelm
parents: 46125
diff changeset
   266
      fun con_strict (con, args) =
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   267
        let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   268
          val rules = abs_strict :: @{thms con_strict_rules}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   269
          val (vs, nonlazy) = get_vars args
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   270
          fun one_strict v' =
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   271
            let
41429
cf5f025bc3c7 renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
huffman
parents: 41296
diff changeset
   272
              val bottom = mk_bottom (fastype_of v')
cf5f025bc3c7 renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
huffman
parents: 41296
diff changeset
   273
              val vs' = map (fn v => if v = v' then bottom else v) vs
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   274
              val goal = mk_trp (mk_undef (list_ccomb (con, vs')))
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   275
              fun tacs ctxt = [simp_tac (put_simpset HOL_basic_ss ctxt addsimps rules) 1]
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   276
            in prove thy con_betas goal (tacs o #context) end
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   277
        in map one_strict nonlazy end
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   278
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   279
      fun con_defin (con, args) =
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   280
        let
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   281
          fun iff_disj (t, []) = HOLogic.mk_not t
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   282
            | iff_disj (t, ts) = mk_eq (t, foldr1 HOLogic.mk_disj ts)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   283
          val (vs, nonlazy) = get_vars args
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   284
          val lhs = mk_undef (list_ccomb (con, vs))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   285
          val rhss = map mk_undef nonlazy
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   286
          val goal = mk_trp (iff_disj (lhs, rhss))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   287
          val rule1 = iso_locale RS @{thm iso.abs_bottom_iff}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   288
          val rules = rule1 :: @{thms con_bottom_iff_rules}
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   289
          fun tacs ctxt = [simp_tac (put_simpset HOL_ss ctxt addsimps rules) 1]
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   290
        in prove thy con_betas goal (tacs o #context) end
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   291
    in
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   292
      val con_stricts = maps con_strict spec'
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   293
      val con_defins = map con_defin spec'
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   294
      val con_rews = con_stricts @ con_defins
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   295
    end
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   296
35456
5356534f880e move proofs of injects and inverts to domain_constructors.ML
huffman
parents: 35455
diff changeset
   297
    (* prove injectiveness of constructors *)
5356534f880e move proofs of injects and inverts to domain_constructors.ML
huffman
parents: 35455
diff changeset
   298
    local
5356534f880e move proofs of injects and inverts to domain_constructors.ML
huffman
parents: 35455
diff changeset
   299
      fun pgterm rel (con, args) =
5356534f880e move proofs of injects and inverts to domain_constructors.ML
huffman
parents: 35455
diff changeset
   300
        let
5356534f880e move proofs of injects and inverts to domain_constructors.ML
huffman
parents: 35455
diff changeset
   301
          fun prime (Free (n, T)) = Free (n^"'", T)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   302
            | prime t             = t
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   303
          val (xs, nonlazy) = get_vars args
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   304
          val ys = map prime xs
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   305
          val lhs = rel (list_ccomb (con, xs), list_ccomb (con, ys))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   306
          val rhs = foldr1 mk_conj (ListPair.map rel (xs, ys))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   307
          val concl = mk_trp (mk_eq (lhs, rhs))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   308
          val zs = case args of [_] => [] | _ => nonlazy
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   309
          val assms = map (mk_trp o mk_defined) zs
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   310
          val goal = Logic.list_implies (assms, concl)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   311
        in prove thy con_betas goal end
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   312
      val cons' = filter (fn (_, args) => not (null args)) spec'
35456
5356534f880e move proofs of injects and inverts to domain_constructors.ML
huffman
parents: 35455
diff changeset
   313
    in
5356534f880e move proofs of injects and inverts to domain_constructors.ML
huffman
parents: 35455
diff changeset
   314
      val inverts =
5356534f880e move proofs of injects and inverts to domain_constructors.ML
huffman
parents: 35455
diff changeset
   315
        let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   316
          val abs_below = iso_locale RS @{thm iso.abs_below}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   317
          val rules1 = abs_below :: @{thms sinl_below sinr_below spair_below up_below}
35456
5356534f880e move proofs of injects and inverts to domain_constructors.ML
huffman
parents: 35455
diff changeset
   318
          val rules2 = @{thms up_defined spair_defined ONE_defined}
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   319
          val rules = rules1 @ rules2
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   320
          fun tacs ctxt = [asm_simp_tac (put_simpset simple_ss ctxt addsimps rules) 1]
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   321
        in map (fn c => pgterm mk_below c (tacs o #context)) cons' end
35456
5356534f880e move proofs of injects and inverts to domain_constructors.ML
huffman
parents: 35455
diff changeset
   322
      val injects =
5356534f880e move proofs of injects and inverts to domain_constructors.ML
huffman
parents: 35455
diff changeset
   323
        let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   324
          val abs_eq = iso_locale RS @{thm iso.abs_eq}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   325
          val rules1 = abs_eq :: @{thms sinl_eq sinr_eq spair_eq up_eq}
35456
5356534f880e move proofs of injects and inverts to domain_constructors.ML
huffman
parents: 35455
diff changeset
   326
          val rules2 = @{thms up_defined spair_defined ONE_defined}
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   327
          val rules = rules1 @ rules2
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   328
          fun tacs ctxt = [asm_simp_tac (put_simpset simple_ss ctxt addsimps rules) 1]
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   329
        in map (fn c => pgterm mk_eq c (tacs o #context)) cons' end
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   330
    end
35456
5356534f880e move proofs of injects and inverts to domain_constructors.ML
huffman
parents: 35455
diff changeset
   331
35458
deaf221c4a59 moved proofs of dist_les and dist_eqs to domain_constructors.ML
huffman
parents: 35457
diff changeset
   332
    (* prove distinctness of constructors *)
deaf221c4a59 moved proofs of dist_les and dist_eqs to domain_constructors.ML
huffman
parents: 35457
diff changeset
   333
    local
deaf221c4a59 moved proofs of dist_les and dist_eqs to domain_constructors.ML
huffman
parents: 35457
diff changeset
   334
      fun map_dist (f : 'a -> 'a -> 'b) (xs : 'a list) : 'b list =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   335
        flat (map_index (fn (i, x) => map (f x) (nth_drop i xs)) xs)
35458
deaf221c4a59 moved proofs of dist_les and dist_eqs to domain_constructors.ML
huffman
parents: 35457
diff changeset
   336
      fun prime (Free (n, T)) = Free (n^"'", T)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   337
        | prime t             = t
35458
deaf221c4a59 moved proofs of dist_les and dist_eqs to domain_constructors.ML
huffman
parents: 35457
diff changeset
   338
      fun iff_disj (t, []) = mk_not t
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   339
        | iff_disj (t, ts) = mk_eq (t, foldr1 mk_disj ts)
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   340
      fun iff_disj2 (t, [], _) = mk_not t
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   341
        | iff_disj2 (t, _, []) = mk_not t
35458
deaf221c4a59 moved proofs of dist_les and dist_eqs to domain_constructors.ML
huffman
parents: 35457
diff changeset
   342
        | iff_disj2 (t, ts, us) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   343
          mk_eq (t, mk_conj (foldr1 mk_disj ts, foldr1 mk_disj us))
35458
deaf221c4a59 moved proofs of dist_les and dist_eqs to domain_constructors.ML
huffman
parents: 35457
diff changeset
   344
      fun dist_le (con1, args1) (con2, args2) =
deaf221c4a59 moved proofs of dist_les and dist_eqs to domain_constructors.ML
huffman
parents: 35457
diff changeset
   345
        let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   346
          val (vs1, zs1) = get_vars args1
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   347
          val (vs2, _) = get_vars args2 |> pairself (map prime)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   348
          val lhs = mk_below (list_ccomb (con1, vs1), list_ccomb (con2, vs2))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   349
          val rhss = map mk_undef zs1
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   350
          val goal = mk_trp (iff_disj (lhs, rhss))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   351
          val rule1 = iso_locale RS @{thm iso.abs_below}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   352
          val rules = rule1 :: @{thms con_below_iff_rules}
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   353
          fun tacs ctxt = [simp_tac (put_simpset HOL_ss ctxt addsimps rules) 1]
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   354
        in prove thy con_betas goal (tacs o #context) end
35458
deaf221c4a59 moved proofs of dist_les and dist_eqs to domain_constructors.ML
huffman
parents: 35457
diff changeset
   355
      fun dist_eq (con1, args1) (con2, args2) =
deaf221c4a59 moved proofs of dist_les and dist_eqs to domain_constructors.ML
huffman
parents: 35457
diff changeset
   356
        let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   357
          val (vs1, zs1) = get_vars args1
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   358
          val (vs2, zs2) = get_vars args2 |> pairself (map prime)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   359
          val lhs = mk_eq (list_ccomb (con1, vs1), list_ccomb (con2, vs2))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   360
          val rhss1 = map mk_undef zs1
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   361
          val rhss2 = map mk_undef zs2
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   362
          val goal = mk_trp (iff_disj2 (lhs, rhss1, rhss2))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   363
          val rule1 = iso_locale RS @{thm iso.abs_eq}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   364
          val rules = rule1 :: @{thms con_eq_iff_rules}
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   365
          fun tacs ctxt = [simp_tac (put_simpset HOL_ss ctxt addsimps rules) 1]
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   366
        in prove thy con_betas goal (tacs o #context) end
35458
deaf221c4a59 moved proofs of dist_les and dist_eqs to domain_constructors.ML
huffman
parents: 35457
diff changeset
   367
    in
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   368
      val dist_les = map_dist dist_le spec'
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   369
      val dist_eqs = map_dist dist_eq spec'
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   370
    end
35456
5356534f880e move proofs of injects and inverts to domain_constructors.ML
huffman
parents: 35455
diff changeset
   371
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   372
    val result =
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   373
      {
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   374
        con_consts = con_consts,
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   375
        con_betas = con_betas,
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   376
        nchotomy = nchotomy,
35457
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   377
        exhaust = exhaust,
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   378
        compacts = compacts,
35456
5356534f880e move proofs of injects and inverts to domain_constructors.ML
huffman
parents: 35455
diff changeset
   379
        con_rews = con_rews,
5356534f880e move proofs of injects and inverts to domain_constructors.ML
huffman
parents: 35455
diff changeset
   380
        inverts = inverts,
35458
deaf221c4a59 moved proofs of dist_les and dist_eqs to domain_constructors.ML
huffman
parents: 35457
diff changeset
   381
        injects = injects,
deaf221c4a59 moved proofs of dist_les and dist_eqs to domain_constructors.ML
huffman
parents: 35457
diff changeset
   382
        dist_les = dist_les,
deaf221c4a59 moved proofs of dist_les and dist_eqs to domain_constructors.ML
huffman
parents: 35457
diff changeset
   383
        dist_eqs = dist_eqs
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   384
      }
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   385
  in
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   386
    (result, thy)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   387
  end
35453
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   388
debbdbb45fbc reorder sections
huffman
parents: 35452
diff changeset
   389
(******************************************************************************)
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   390
(**************** definition and theorems for case combinator *****************)
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   391
(******************************************************************************)
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   392
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   393
fun add_case_combinator
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   394
    (spec : (term * (bool * typ) list) list)
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   395
    (lhsT : typ)
35777
bcc77916b7b9 pass binding as argument to add_domain_constructors; proper binding for case combinator
huffman
parents: 35561
diff changeset
   396
    (dbind : binding)
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   397
    (con_betas : thm list)
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   398
    (iso_locale : thm)
35486
c91854705b1d move definition of case combinator to domain_constructors.ML
huffman
parents: 35485
diff changeset
   399
    (rep_const : term)
35460
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   400
    (thy : theory)
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   401
    : ((typ -> term) * thm list) * theory =
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   402
  let
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   403
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   404
    (* prove rep/abs rules *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   405
    val rep_strict = iso_locale RS @{thm iso.rep_strict}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   406
    val abs_inverse = iso_locale RS @{thm iso.abs_iso}
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   407
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   408
    (* calculate function arguments of case combinator *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   409
    val tns = map fst (Term.add_tfreesT lhsT [])
43324
2b47822868e4 discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents: 42290
diff changeset
   410
    val resultT = TFree (singleton (Name.variant_list tns) "'t", @{sort pcpo})
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   411
    fun fTs T = map (fn (_, args) => map snd args -->> T) spec
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 54895
diff changeset
   412
    val fns = Old_Datatype_Prop.indexify_names (map (K "f") spec)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   413
    val fs = map Free (fns ~~ fTs resultT)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   414
    fun caseT T = fTs T -->> (lhsT ->> T)
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   415
35486
c91854705b1d move definition of case combinator to domain_constructors.ML
huffman
parents: 35485
diff changeset
   416
    (* definition of case combinator *)
c91854705b1d move definition of case combinator to domain_constructors.ML
huffman
parents: 35485
diff changeset
   417
    local
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   418
      val case_bind = Binding.suffix_name "_case" dbind
35784
a86ed293b005 simplify definition of when combinators
huffman
parents: 35781
diff changeset
   419
      fun lambda_arg (lazy, v) t =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   420
          (if lazy then mk_fup else I) (big_lambda v t)
40212
20df78048db5 rename constant 'one_when' to 'one_case'
huffman
parents: 40019
diff changeset
   421
      fun lambda_args []      t = mk_one_case t
35784
a86ed293b005 simplify definition of when combinators
huffman
parents: 35781
diff changeset
   422
        | lambda_args (x::[]) t = lambda_arg x t
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   423
        | lambda_args (x::xs) t = mk_ssplit (lambda_arg x (lambda_args xs t))
35486
c91854705b1d move definition of case combinator to domain_constructors.ML
huffman
parents: 35485
diff changeset
   424
      fun one_con f (_, args) =
c91854705b1d move definition of case combinator to domain_constructors.ML
huffman
parents: 35485
diff changeset
   425
        let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   426
          val Ts = map snd args
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 54895
diff changeset
   427
          val ns = Name.variant_list fns (Old_Datatype_Prop.make_tnames Ts)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   428
          val vs = map Free (ns ~~ Ts)
35486
c91854705b1d move definition of case combinator to domain_constructors.ML
huffman
parents: 35485
diff changeset
   429
        in
35784
a86ed293b005 simplify definition of when combinators
huffman
parents: 35781
diff changeset
   430
          lambda_args (map fst args ~~ vs) (list_ccomb (f, vs))
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   431
        end
35785
cdaf99fddd17 move functions into holcf_library.ML
huffman
parents: 35784
diff changeset
   432
      fun mk_sscases [t] = mk_strictify t
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   433
        | mk_sscases ts = foldr1 mk_sscase ts
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   434
      val body = mk_sscases (map2 one_con fs spec)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   435
      val rhs = big_lambdas fs (mk_cfcomp (body, rep_const))
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   436
      val ((_, case_defs), thy) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   437
          define_consts [(case_bind, rhs, NoSyn)] thy
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   438
      val case_name = Sign.full_name thy case_bind
35486
c91854705b1d move definition of case combinator to domain_constructors.ML
huffman
parents: 35485
diff changeset
   439
    in
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   440
      val case_def = hd case_defs
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   441
      fun case_const T = Const (case_name, caseT T)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   442
      val case_app = list_ccomb (case_const resultT, fs)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   443
      val thy = thy
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   444
    end
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   445
35472
c23b42730b9b move case combinator syntax to domain_constructors.ML
huffman
parents: 35470
diff changeset
   446
    (* define syntax for case combinator *)
c23b42730b9b move case combinator syntax to domain_constructors.ML
huffman
parents: 35470
diff changeset
   447
    (* TODO: re-implement case syntax using a parse translation *)
c23b42730b9b move case combinator syntax to domain_constructors.ML
huffman
parents: 35470
diff changeset
   448
    local
42290
b1f544c84040 discontinued special treatment of structure Lexicon;
wenzelm
parents: 42264
diff changeset
   449
      fun syntax c = Lexicon.mark_const (fst (dest_Const c))
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   450
      fun xconst c = Long_Name.base_name (fst (dest_Const c))
42224
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42204
diff changeset
   451
      fun c_ast authentic con = Ast.Constant (if authentic then syntax con else xconst con)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   452
      fun showint n = string_of_int (n+1)
42224
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42204
diff changeset
   453
      fun expvar n = Ast.Variable ("e" ^ showint n)
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42204
diff changeset
   454
      fun argvar n (m, _) = Ast.Variable ("a" ^ showint n ^ "_" ^ showint m)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   455
      fun argvars n args = map_index (argvar n) args
42224
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42204
diff changeset
   456
      fun app s (l, r) = Ast.mk_appl (Ast.Constant s) [l, r]
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   457
      val cabs = app "_cabs"
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   458
      val capp = app @{const_syntax Rep_cfun}
35472
c23b42730b9b move case combinator syntax to domain_constructors.ML
huffman
parents: 35470
diff changeset
   459
      val capps = Library.foldl capp
46125
00cd193a48dc improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents: 45654
diff changeset
   460
      fun con1 authentic n (con, args) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   461
          Library.foldl capp (c_ast authentic con, argvars n args)
46125
00cd193a48dc improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents: 45654
diff changeset
   462
      fun con1_constraint authentic n (con, args) =
00cd193a48dc improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents: 45654
diff changeset
   463
          Library.foldl capp
00cd193a48dc improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents: 45654
diff changeset
   464
            (Ast.Appl
00cd193a48dc improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents: 45654
diff changeset
   465
              [Ast.Constant @{syntax_const "_constrain"}, c_ast authentic con,
00cd193a48dc improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents: 45654
diff changeset
   466
                Ast.Variable ("'a" ^ string_of_int n)],
00cd193a48dc improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents: 45654
diff changeset
   467
             argvars n args)
00cd193a48dc improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents: 45654
diff changeset
   468
      fun case1 constraint authentic (n, c) =
00cd193a48dc improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents: 45654
diff changeset
   469
        app @{syntax_const "_case1"}
00cd193a48dc improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents: 45654
diff changeset
   470
          ((if constraint then con1_constraint else con1) authentic n c, expvar n)
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   471
      fun arg1 (n, (_, args)) = List.foldr cabs (expvar n) (argvars n args)
42224
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42204
diff changeset
   472
      fun when1 n (m, c) = if n = m then arg1 (n, c) else Ast.Constant @{const_syntax bottom}
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42204
diff changeset
   473
      val case_constant = Ast.Constant (syntax (case_const dummyT))
46125
00cd193a48dc improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents: 45654
diff changeset
   474
      fun case_trans constraint authentic =
42224
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42204
diff changeset
   475
          (app "_case_syntax"
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42204
diff changeset
   476
            (Ast.Variable "x",
46125
00cd193a48dc improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents: 45654
diff changeset
   477
             foldr1 (app @{syntax_const "_case2"}) (map_index (case1 constraint authentic) spec)),
42224
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42204
diff changeset
   478
           capp (capps (case_constant, map_index arg1 spec), Ast.Variable "x"))
35485
7d7495f5e35e remove dependence on Domain_Library
huffman
parents: 35484
diff changeset
   479
      fun one_abscon_trans authentic (n, c) =
46125
00cd193a48dc improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents: 45654
diff changeset
   480
          (if authentic then Syntax.Parse_Print_Rule else Syntax.Parse_Rule)
00cd193a48dc improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents: 45654
diff changeset
   481
            (cabs (con1 authentic n c, expvar n),
00cd193a48dc improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents: 45654
diff changeset
   482
             capps (case_constant, map_index (when1 n) spec))
35472
c23b42730b9b move case combinator syntax to domain_constructors.ML
huffman
parents: 35470
diff changeset
   483
      fun abscon_trans authentic =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   484
          map_index (one_abscon_trans authentic) spec
42224
578a51fae383 discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
wenzelm
parents: 42204
diff changeset
   485
      val trans_rules : Ast.ast Syntax.trrule list =
46125
00cd193a48dc improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents: 45654
diff changeset
   486
          Syntax.Parse_Print_Rule (case_trans false true) ::
00cd193a48dc improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents: 45654
diff changeset
   487
          Syntax.Parse_Rule (case_trans false false) ::
00cd193a48dc improved case syntax: more careful treatment of position constraints, which enables PIDE markup;
wenzelm
parents: 45654
diff changeset
   488
          Syntax.Parse_Rule (case_trans true false) ::
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   489
          abscon_trans false @ abscon_trans true
35472
c23b42730b9b move case combinator syntax to domain_constructors.ML
huffman
parents: 35470
diff changeset
   490
    in
42204
b3277168c1e7 added Position.reports convenience;
wenzelm
parents: 42151
diff changeset
   491
      val thy = Sign.add_trrules trans_rules thy
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   492
    end
35472
c23b42730b9b move case combinator syntax to domain_constructors.ML
huffman
parents: 35470
diff changeset
   493
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   494
    (* prove beta reduction rule for case combinator *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   495
    val case_beta = beta_of_def thy case_def
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   496
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   497
    (* prove strictness of case combinator *)
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   498
    val case_strict =
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   499
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   500
        val defs = case_beta :: map mk_meta_eq [rep_strict, @{thm cfcomp2}]
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   501
        val goal = mk_trp (mk_strict case_app)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   502
        val rules = @{thms sscase1 ssplit1 strictify1 one_case1}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   503
        val tacs = [resolve_tac rules 1]
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   504
      in prove thy defs goal (K tacs) end
46490
e4863ab5e09b more basic simplification, eliminated slightly odd tactic style from 1995 (cf. ea0668a1c0ba);
wenzelm
parents: 46125
diff changeset
   505
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   506
    (* prove rewrites for case combinator *)
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   507
    local
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   508
      fun one_case (con, args) f =
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   509
        let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   510
          val (vs, nonlazy) = get_vars args
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   511
          val assms = map (mk_trp o mk_defined) nonlazy
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   512
          val lhs = case_app ` list_ccomb (con, vs)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   513
          val rhs = list_ccomb (f, vs)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   514
          val concl = mk_trp (mk_eq (lhs, rhs))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   515
          val goal = Logic.list_implies (assms, concl)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   516
          val defs = case_beta :: con_betas
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   517
          val rules1 = @{thms strictify2 sscase2 sscase3 ssplit2 fup2 ID1}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   518
          val rules2 = @{thms con_bottom_iff_rules}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   519
          val rules3 = @{thms cfcomp2 one_case2}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   520
          val rules = abs_inverse :: rules1 @ rules2 @ rules3
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   521
          fun tacs ctxt = [asm_simp_tac (put_simpset beta_ss ctxt addsimps rules) 1]
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   522
        in prove thy defs goal (tacs o #context) end
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   523
    in
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   524
      val case_apps = map2 one_case spec fs
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   525
    end
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   526
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   527
  in
35460
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   528
    ((case_const, case_strict :: case_apps), thy)
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   529
  end
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   530
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   531
(******************************************************************************)
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   532
(************** definitions and theorems for selector functions ***************)
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   533
(******************************************************************************)
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   534
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   535
fun add_selectors
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   536
    (spec : (term * (bool * binding option * typ) list) list)
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   537
    (rep_const : term)
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   538
    (abs_inv : thm)
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   539
    (rep_strict : thm)
40321
d065b195ec89 rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
huffman
parents: 40214
diff changeset
   540
    (rep_bottom_iff : thm)
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   541
    (con_betas : thm list)
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   542
    (thy : theory)
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   543
    : thm list * theory =
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   544
  let
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   545
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   546
    (* define selector functions *)
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   547
    val ((sel_consts, sel_defs), thy) =
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   548
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   549
        fun rangeT s = snd (dest_cfunT (fastype_of s))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   550
        fun mk_outl s = mk_cfcomp (from_sinl (dest_ssumT (rangeT s)), s)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   551
        fun mk_outr s = mk_cfcomp (from_sinr (dest_ssumT (rangeT s)), s)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   552
        fun mk_sfst s = mk_cfcomp (sfst_const (dest_sprodT (rangeT s)), s)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   553
        fun mk_ssnd s = mk_cfcomp (ssnd_const (dest_sprodT (rangeT s)), s)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   554
        fun mk_down s = mk_cfcomp (from_up (dest_upT (rangeT s)), s)
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   555
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   556
        fun sels_of_arg _ (_, NONE, _) = []
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   557
          | sels_of_arg s (lazy, SOME b, _) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   558
            [(b, if lazy then mk_down s else s, NoSyn)]
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   559
        fun sels_of_args _ [] = []
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   560
          | sels_of_args s (v :: []) = sels_of_arg s v
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   561
          | sels_of_args s (v :: vs) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   562
            sels_of_arg (mk_sfst s) v @ sels_of_args (mk_ssnd s) vs
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   563
        fun sels_of_cons _ [] = []
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   564
          | sels_of_cons s ((_, args) :: []) = sels_of_args s args
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   565
          | sels_of_cons s ((_, args) :: cs) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   566
            sels_of_args (mk_outl s) args @ sels_of_cons (mk_outr s) cs
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   567
        val sel_eqns : (binding * term * mixfix) list =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   568
            sels_of_cons rep_const spec
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   569
      in
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   570
        define_consts sel_eqns thy
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   571
      end
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   572
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   573
    (* replace bindings with terms in constructor spec *)
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   574
    val spec2 : (term * (bool * term option * typ) list) list =
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   575
      let
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   576
        fun prep_arg (lazy, NONE, T) sels = ((lazy, NONE, T), sels)
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   577
          | prep_arg (lazy, SOME _, T) sels =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   578
            ((lazy, SOME (hd sels), T), tl sels)
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   579
        fun prep_con (con, args) sels =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   580
            apfst (pair con) (fold_map prep_arg args sels)
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   581
      in
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   582
        fst (fold_map prep_con spec sel_consts)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   583
      end
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   584
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   585
    (* prove selector strictness rules *)
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   586
    val sel_stricts : thm list =
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   587
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   588
        val rules = rep_strict :: @{thms sel_strict_rules}
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   589
        fun tacs ctxt = [simp_tac (put_simpset HOL_basic_ss ctxt addsimps rules) 1]
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   590
        fun sel_strict sel =
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   591
          let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   592
            val goal = mk_trp (mk_strict sel)
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   593
          in
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   594
            prove thy sel_defs goal (tacs o #context)
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   595
          end
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   596
      in
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   597
        map sel_strict sel_consts
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   598
      end
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   599
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   600
    (* prove selector application rules *)
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   601
    val sel_apps : thm list =
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   602
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   603
        val defs = con_betas @ sel_defs
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   604
        val rules = abs_inv :: @{thms sel_app_rules}
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   605
        fun tacs ctxt = [asm_simp_tac (put_simpset simple_ss ctxt addsimps rules) 1]
37165
c2e27ae53c2a made SML/NJ quite happy;
wenzelm
parents: 37109
diff changeset
   606
        fun sel_apps_of (i, (con, args: (bool * term option * typ) list)) =
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   607
          let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   608
            val Ts : typ list = map #3 args
58112
8081087096ad renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
blanchet
parents: 54895
diff changeset
   609
            val ns : string list = Old_Datatype_Prop.make_tnames Ts
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   610
            val vs : term list = map Free (ns ~~ Ts)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   611
            val con_app : term = list_ccomb (con, vs)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   612
            val vs' : (bool * term) list = map #1 args ~~ vs
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   613
            fun one_same (n, sel, _) =
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   614
              let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   615
                val xs = map snd (filter_out fst (nth_drop n vs'))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   616
                val assms = map (mk_trp o mk_defined) xs
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   617
                val concl = mk_trp (mk_eq (sel ` con_app, nth vs n))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   618
                val goal = Logic.list_implies (assms, concl)
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   619
              in
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   620
                prove thy defs goal (tacs o #context)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   621
              end
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   622
            fun one_diff (_, sel, T) =
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   623
              let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   624
                val goal = mk_trp (mk_eq (sel ` con_app, mk_bottom T))
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   625
              in
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   626
                prove thy defs goal (tacs o #context)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   627
              end
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   628
            fun one_con (j, (_, args')) : thm list =
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   629
              let
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   630
                fun prep (_, (_, NONE, _)) = NONE
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   631
                  | prep (i, (_, SOME sel, T)) = SOME (i, sel, T)
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   632
                val sels : (int * term * typ) list =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   633
                  map_filter prep (map_index I args')
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   634
              in
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   635
                if i = j
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   636
                then map one_same sels
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   637
                else map one_diff sels
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   638
              end
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   639
          in
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   640
            flat (map_index one_con spec2)
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   641
          end
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   642
      in
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   643
        flat (map_index sel_apps_of spec2)
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   644
      end
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   645
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   646
  (* prove selector definedness rules *)
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   647
    val sel_defins : thm list =
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   648
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   649
        val rules = rep_bottom_iff :: @{thms sel_bottom_iff_rules}
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   650
        fun tacs ctxt = [simp_tac (put_simpset HOL_basic_ss ctxt addsimps rules) 1]
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   651
        fun sel_defin sel =
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   652
          let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   653
            val (T, U) = dest_cfunT (fastype_of sel)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   654
            val x = Free ("x", T)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   655
            val lhs = mk_eq (sel ` x, mk_bottom U)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   656
            val rhs = mk_eq (x, mk_bottom T)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   657
            val goal = mk_trp (mk_eq (lhs, rhs))
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   658
          in
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   659
            prove thy sel_defs goal (tacs o #context)
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   660
          end
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   661
        fun one_arg (false, SOME sel, _) = SOME (sel_defin sel)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   662
          | one_arg _                    = NONE
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   663
      in
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   664
        case spec2 of
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   665
          [(_, args)] => map_filter one_arg args
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   666
        | _           => []
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   667
      end
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   668
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   669
  in
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   670
    (sel_stricts @ sel_defins @ sel_apps, thy)
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   671
  end
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   672
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   673
(******************************************************************************)
35460
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   674
(************ definitions and theorems for discriminator functions ************)
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   675
(******************************************************************************)
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   676
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   677
fun add_discriminators
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   678
    (bindings : binding list)
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   679
    (spec : (term * (bool * typ) list) list)
35461
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   680
    (lhsT : typ)
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   681
    (exhaust : thm)
35460
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   682
    (case_const : typ -> term)
35461
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   683
    (case_rews : thm list)
35460
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   684
    (thy : theory) =
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   685
  let
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   686
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   687
    (* define discriminator functions *)
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   688
    local
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   689
      fun dis_fun i (j, (_, args)) =
35460
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   690
        let
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   691
          val (vs, _) = get_vars args
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   692
          val tr = if i = j then @{term TT} else @{term FF}
35460
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   693
        in
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   694
          big_lambdas vs tr
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   695
        end
35460
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   696
      fun dis_eqn (i, bind) : binding * term * mixfix =
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   697
        let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   698
          val dis_bind = Binding.prefix_name "is_" bind
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   699
          val rhs = list_ccomb (case_const trT, map_index (dis_fun i) spec)
35460
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   700
        in
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   701
          (dis_bind, rhs, NoSyn)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   702
        end
35460
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   703
    in
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   704
      val ((dis_consts, dis_defs), thy) =
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   705
          define_consts (map_index dis_eqn bindings) thy
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   706
    end
35460
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   707
35461
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   708
    (* prove discriminator strictness rules *)
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   709
    local
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   710
      fun dis_strict dis =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   711
        let val goal = mk_trp (mk_strict dis)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   712
        in prove thy dis_defs goal (K [rtac (hd case_rews) 1]) end
35461
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   713
    in
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   714
      val dis_stricts = map dis_strict dis_consts
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   715
    end
35461
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   716
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   717
    (* prove discriminator/constructor rules *)
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   718
    local
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   719
      fun dis_app (i, dis) (j, (con, args)) =
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   720
        let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   721
          val (vs, nonlazy) = get_vars args
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   722
          val lhs = dis ` list_ccomb (con, vs)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   723
          val rhs = if i = j then @{term TT} else @{term FF}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   724
          val assms = map (mk_trp o mk_defined) nonlazy
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   725
          val concl = mk_trp (mk_eq (lhs, rhs))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   726
          val goal = Logic.list_implies (assms, concl)
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   727
          fun tacs ctxt = [asm_simp_tac (put_simpset beta_ss ctxt addsimps case_rews) 1]
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   728
        in prove thy dis_defs goal (tacs o #context) end
35461
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   729
      fun one_dis (i, dis) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   730
          map_index (dis_app (i, dis)) spec
35461
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   731
    in
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   732
      val dis_apps = flat (map_index one_dis dis_consts)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   733
    end
35461
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   734
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   735
    (* prove discriminator definedness rules *)
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   736
    local
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   737
      fun dis_defin dis =
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   738
        let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   739
          val x = Free ("x", lhsT)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   740
          val simps = dis_apps @ @{thms dist_eq_tr}
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   741
          fun tacs ctxt =
35461
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   742
            [rtac @{thm iffI} 1,
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   743
             asm_simp_tac (put_simpset HOL_basic_ss ctxt addsimps dis_stricts) 2,
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   744
             rtac exhaust 1, atac 1,
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   745
             ALLGOALS (asm_full_simp_tac (put_simpset simple_ss ctxt addsimps simps))]
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   746
          val goal = mk_trp (mk_eq (mk_undef (dis ` x), mk_undef x))
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   747
        in prove thy [] goal (tacs o #context) end
35461
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   748
    in
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   749
      val dis_defins = map dis_defin dis_consts
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   750
    end
35461
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   751
35460
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   752
  in
35461
34360a1e3537 move proofs of dis_rews to domain_constructors.ML; change dis_defins to iff-style
huffman
parents: 35460
diff changeset
   753
    (dis_stricts @ dis_defins @ dis_apps, thy)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   754
  end
35460
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   755
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   756
(******************************************************************************)
35462
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   757
(*************** definitions and theorems for match combinators ***************)
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   758
(******************************************************************************)
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   759
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   760
fun add_match_combinators
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   761
    (bindings : binding list)
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   762
    (spec : (term * (bool * typ) list) list)
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   763
    (lhsT : typ)
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   764
    (case_const : typ -> term)
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   765
    (case_rews : thm list)
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   766
    (thy : theory) =
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   767
  let
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   768
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   769
    (* get a fresh type variable for the result type *)
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   770
    val resultT : typ =
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   771
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   772
        val ts : string list = map fst (Term.add_tfreesT lhsT [])
43324
2b47822868e4 discontinued Name.variant to emphasize that this is old-style / indirect;
wenzelm
parents: 42290
diff changeset
   773
        val t : string = singleton (Name.variant_list ts) "'t"
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   774
      in TFree (t, @{sort pcpo}) end
35462
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   775
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   776
    (* define match combinators *)
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   777
    local
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   778
      val x = Free ("x", lhsT)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   779
      fun k args = Free ("k", map snd args -->> mk_matchT resultT)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   780
      val fail = mk_fail resultT
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   781
      fun mat_fun i (j, (_, args)) =
35462
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   782
        let
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   783
          val (vs, _) = get_vars_avoiding ["x","k"] args
35462
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   784
        in
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   785
          if i = j then k args else big_lambdas vs fail
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   786
        end
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   787
      fun mat_eqn (i, (bind, (_, args))) : binding * term * mixfix =
35462
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   788
        let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   789
          val mat_bind = Binding.prefix_name "match_" bind
35462
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   790
          val funs = map_index (mat_fun i) spec
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   791
          val body = list_ccomb (case_const (mk_matchT resultT), funs)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   792
          val rhs = big_lambda x (big_lambda (k args) (body ` x))
35462
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   793
        in
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   794
          (mat_bind, rhs, NoSyn)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   795
        end
35462
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   796
    in
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   797
      val ((match_consts, match_defs), thy) =
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   798
          define_consts (map_index mat_eqn (bindings ~~ spec)) thy
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   799
    end
35462
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   800
35463
b20501588930 register match functions from domain_constructors.ML
huffman
parents: 35462
diff changeset
   801
    (* register match combinators with fixrec package *)
b20501588930 register match functions from domain_constructors.ML
huffman
parents: 35462
diff changeset
   802
    local
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   803
      val con_names = map (fst o dest_Const o fst) spec
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   804
      val mat_names = map (fst o dest_Const) match_consts
35463
b20501588930 register match functions from domain_constructors.ML
huffman
parents: 35462
diff changeset
   805
    in
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   806
      val thy = Fixrec.add_matchers (con_names ~~ mat_names) thy
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   807
    end
35463
b20501588930 register match functions from domain_constructors.ML
huffman
parents: 35462
diff changeset
   808
35466
9fcfd5763181 move proofs of match_rews to domain_constructors.ML
huffman
parents: 35463
diff changeset
   809
    (* prove strictness of match combinators *)
9fcfd5763181 move proofs of match_rews to domain_constructors.ML
huffman
parents: 35463
diff changeset
   810
    local
9fcfd5763181 move proofs of match_rews to domain_constructors.ML
huffman
parents: 35463
diff changeset
   811
      fun match_strict mat =
9fcfd5763181 move proofs of match_rews to domain_constructors.ML
huffman
parents: 35463
diff changeset
   812
        let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   813
          val (T, (U, V)) = apsnd dest_cfunT (dest_cfunT (fastype_of mat))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   814
          val k = Free ("k", U)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   815
          val goal = mk_trp (mk_eq (mat ` mk_bottom T ` k, mk_bottom V))
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   816
          fun tacs ctxt = [asm_simp_tac (put_simpset beta_ss ctxt addsimps case_rews) 1]
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   817
        in prove thy match_defs goal (tacs o #context) end
35466
9fcfd5763181 move proofs of match_rews to domain_constructors.ML
huffman
parents: 35463
diff changeset
   818
    in
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   819
      val match_stricts = map match_strict match_consts
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   820
    end
35466
9fcfd5763181 move proofs of match_rews to domain_constructors.ML
huffman
parents: 35463
diff changeset
   821
9fcfd5763181 move proofs of match_rews to domain_constructors.ML
huffman
parents: 35463
diff changeset
   822
    (* prove match/constructor rules *)
9fcfd5763181 move proofs of match_rews to domain_constructors.ML
huffman
parents: 35463
diff changeset
   823
    local
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   824
      val fail = mk_fail resultT
35466
9fcfd5763181 move proofs of match_rews to domain_constructors.ML
huffman
parents: 35463
diff changeset
   825
      fun match_app (i, mat) (j, (con, args)) =
9fcfd5763181 move proofs of match_rews to domain_constructors.ML
huffman
parents: 35463
diff changeset
   826
        let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   827
          val (vs, nonlazy) = get_vars_avoiding ["k"] args
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   828
          val (_, (kT, _)) = apsnd dest_cfunT (dest_cfunT (fastype_of mat))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   829
          val k = Free ("k", kT)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   830
          val lhs = mat ` list_ccomb (con, vs) ` k
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   831
          val rhs = if i = j then list_ccomb (k, vs) else fail
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   832
          val assms = map (mk_trp o mk_defined) nonlazy
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   833
          val concl = mk_trp (mk_eq (lhs, rhs))
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   834
          val goal = Logic.list_implies (assms, concl)
54895
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   835
          fun tacs ctxt = [asm_simp_tac (put_simpset beta_ss ctxt addsimps case_rews) 1]
515630483010 clarified simplifier context;
wenzelm
parents: 54742
diff changeset
   836
        in prove thy match_defs goal (tacs o #context) end
35466
9fcfd5763181 move proofs of match_rews to domain_constructors.ML
huffman
parents: 35463
diff changeset
   837
      fun one_match (i, mat) =
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   838
          map_index (match_app (i, mat)) spec
35466
9fcfd5763181 move proofs of match_rews to domain_constructors.ML
huffman
parents: 35463
diff changeset
   839
    in
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   840
      val match_apps = flat (map_index one_match match_consts)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   841
    end
35466
9fcfd5763181 move proofs of match_rews to domain_constructors.ML
huffman
parents: 35463
diff changeset
   842
35462
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   843
  in
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   844
    (match_stricts @ match_apps, thy)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   845
  end
35462
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   846
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   847
(******************************************************************************)
35450
e9ef2b50ac59 move constructor-specific stuff to a separate function
huffman
parents: 35449
diff changeset
   848
(******************************* main function ********************************)
e9ef2b50ac59 move constructor-specific stuff to a separate function
huffman
parents: 35449
diff changeset
   849
(******************************************************************************)
e9ef2b50ac59 move constructor-specific stuff to a separate function
huffman
parents: 35449
diff changeset
   850
e9ef2b50ac59 move constructor-specific stuff to a separate function
huffman
parents: 35449
diff changeset
   851
fun add_domain_constructors
35777
bcc77916b7b9 pass binding as argument to add_domain_constructors; proper binding for case combinator
huffman
parents: 35561
diff changeset
   852
    (dbind : binding)
35481
7bb9157507a9 add_domain_constructors takes iso_info record as argument
huffman
parents: 35476
diff changeset
   853
    (spec : (binding * (bool * binding option * typ) list * mixfix) list)
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents: 35496
diff changeset
   854
    (iso_info : Domain_Take_Proofs.iso_info)
35450
e9ef2b50ac59 move constructor-specific stuff to a separate function
huffman
parents: 35449
diff changeset
   855
    (thy : theory) =
e9ef2b50ac59 move constructor-specific stuff to a separate function
huffman
parents: 35449
diff changeset
   856
  let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   857
    val dname = Binding.name_of dbind
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   858
    val _ = writeln ("Proving isomorphism properties of domain "^dname^" ...")
35450
e9ef2b50ac59 move constructor-specific stuff to a separate function
huffman
parents: 35449
diff changeset
   859
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   860
    val bindings = map #1 spec
40019
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40017
diff changeset
   861
35481
7bb9157507a9 add_domain_constructors takes iso_info record as argument
huffman
parents: 35476
diff changeset
   862
    (* retrieve facts about rep/abs *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   863
    val lhsT = #absT iso_info
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   864
    val {rep_const, abs_const, ...} = iso_info
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   865
    val abs_iso_thm = #abs_inverse iso_info
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   866
    val rep_iso_thm = #rep_inverse iso_info
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   867
    val iso_locale = @{thm iso.intro} OF [abs_iso_thm, rep_iso_thm]
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   868
    val rep_strict = iso_locale RS @{thm iso.rep_strict}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   869
    val abs_strict = iso_locale RS @{thm iso.abs_strict}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   870
    val rep_bottom_iff = iso_locale RS @{thm iso.rep_bottom_iff}
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   871
    val iso_rews = [abs_iso_thm, rep_iso_thm, abs_strict, rep_strict]
35450
e9ef2b50ac59 move constructor-specific stuff to a separate function
huffman
parents: 35449
diff changeset
   872
35487
d1630f317ed0 qualify constructor names with type name
huffman
parents: 35486
diff changeset
   873
    (* qualify constants and theorems with domain name *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   874
    val thy = Sign.add_path dname thy
35487
d1630f317ed0 qualify constructor names with type name
huffman
parents: 35486
diff changeset
   875
35450
e9ef2b50ac59 move constructor-specific stuff to a separate function
huffman
parents: 35449
diff changeset
   876
    (* define constructor functions *)
35457
d63655b88369 move proofs of casedist into domain_constructors.ML
huffman
parents: 35456
diff changeset
   877
    val (con_result, thy) =
35454
cf6ba350b7ca remove unnecessary stuff from argument to add_constructors function
huffman
parents: 35453
diff changeset
   878
      let
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   879
        fun prep_arg (lazy, _, T) = (lazy, T)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   880
        fun prep_con (b, args, mx) = (b, map prep_arg args, mx)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   881
        val con_spec = map prep_con spec
35454
cf6ba350b7ca remove unnecessary stuff from argument to add_constructors function
huffman
parents: 35453
diff changeset
   882
      in
cf6ba350b7ca remove unnecessary stuff from argument to add_constructors function
huffman
parents: 35453
diff changeset
   883
        add_constructors con_spec abs_const iso_locale thy
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   884
      end
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   885
    val {con_consts, con_betas, nchotomy, exhaust, compacts, con_rews,
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   886
          inverts, injects, dist_les, dist_eqs} = con_result
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   887
40019
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40017
diff changeset
   888
    (* prepare constructor spec *)
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40017
diff changeset
   889
    val con_specs : (term * (bool * typ) list) list =
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   890
      let
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   891
        fun prep_arg (lazy, _, T) = (lazy, T)
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   892
        fun prep_con c (_, args, _) = (c, map prep_arg args)
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   893
      in
40019
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40017
diff changeset
   894
        map2 prep_con con_consts spec
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   895
      end
40019
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40017
diff changeset
   896
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40017
diff changeset
   897
    (* define case combinator *)
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40017
diff changeset
   898
    val ((case_const : typ -> term, cases : thm list), thy) =
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40017
diff changeset
   899
        add_case_combinator con_specs lhsT dbind
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   900
          con_betas iso_locale rep_const thy
35450
e9ef2b50ac59 move constructor-specific stuff to a separate function
huffman
parents: 35449
diff changeset
   901
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   902
    (* define and prove theorems for selector functions *)
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   903
    val (sel_thms : thm list, thy : theory) =
35460
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   904
      let
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   905
        val sel_spec : (term * (bool * binding option * typ) list) list =
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   906
          map2 (fn con => fn (_, args, _) => (con, args)) con_consts spec
35460
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   907
      in
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   908
        add_selectors sel_spec rep_const
40321
d065b195ec89 rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
huffman
parents: 40214
diff changeset
   909
          abs_iso_thm rep_strict rep_bottom_iff con_betas thy
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   910
      end
35460
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   911
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   912
    (* define and prove theorems for discriminator functions *)
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   913
    val (dis_thms : thm list, thy : theory) =
40019
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40017
diff changeset
   914
        add_discriminators bindings con_specs lhsT
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   915
          exhaust case_const cases thy
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   916
35462
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   917
    (* define and prove theorems for match combinators *)
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   918
    val (match_thms : thm list, thy : theory) =
40019
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40017
diff changeset
   919
        add_match_combinators bindings con_specs lhsT
44080
53d95b52954c HOLCF: fix warnings about unreferenced identifiers
huffman
parents: 43324
diff changeset
   920
          case_const cases thy
35462
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   921
35446
b719dad322fa rewrite domain package code for selector functions
huffman
parents: 35445
diff changeset
   922
    (* restore original signature path *)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   923
    val thy = Sign.parent_path thy
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
   924
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   925
    (* bind theorem names in global theory *)
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   926
    val (_, thy) =
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   927
      let
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   928
        fun qualified name = Binding.qualified true name dbind
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   929
        val names = "bottom" :: map (fn (b,_,_) => Binding.name_of b) spec
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   930
        val dname = fst (dest_Type lhsT)
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   931
        val simp = Simplifier.simp_add
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   932
        val case_names = Rule_Cases.case_names names
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   933
        val cases_type = Induct.cases_type dname
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   934
      in
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   935
        Global_Theory.add_thmss [
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   936
          ((qualified "iso_rews"  , iso_rews    ), [simp]),
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   937
          ((qualified "nchotomy"  , [nchotomy]  ), []),
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   938
          ((qualified "exhaust"   , [exhaust]   ), [case_names, cases_type]),
40213
b63e966564da rename case combinators generated by domain package to 'foo_case' instead of 'foo_when'
huffman
parents: 40212
diff changeset
   939
          ((qualified "case_rews" , cases       ), [simp]),
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   940
          ((qualified "compacts"  , compacts    ), [simp]),
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   941
          ((qualified "con_rews"  , con_rews    ), [simp]),
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   942
          ((qualified "sel_rews"  , sel_thms    ), [simp]),
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   943
          ((qualified "dis_rews"  , dis_thms    ), [simp]),
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   944
          ((qualified "dist_les"  , dist_les    ), [simp]),
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   945
          ((qualified "dist_eqs"  , dist_eqs    ), [simp]),
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   946
          ((qualified "inverts"   , inverts     ), [simp]),
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   947
          ((qualified "injects"   , injects     ), [simp]),
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   948
          ((qualified "match_rews", match_thms  ), [simp])] thy
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   949
      end
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   950
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
   951
    val result =
40017
575d3aa1f3c5 include iso_info as part of constr_info type
huffman
parents: 40016
diff changeset
   952
      {
575d3aa1f3c5 include iso_info as part of constr_info type
huffman
parents: 40016
diff changeset
   953
        iso_info = iso_info,
40019
05cda34d36e7 put constructor argument specs in constr_info type
huffman
parents: 40017
diff changeset
   954
        con_specs = con_specs,
35451
a726a033b313 don't bother returning con_defs
huffman
parents: 35450
diff changeset
   955
        con_betas = con_betas,
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   956
        nchotomy = nchotomy,
35781
b7738ab762b1 renamed some lemmas generated by the domain package
huffman
parents: 35777
diff changeset
   957
        exhaust = exhaust,
40016
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   958
        compacts = compacts,
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   959
        con_rews = con_rews,
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   960
        inverts = inverts,
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   961
        injects = injects,
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   962
        dist_les = dist_les,
2eff1cbc1ccb remove function Domain_Theorems.theorems; bind theorem names directly from Domain_Constructors.add_domain_constructors
huffman
parents: 40014
diff changeset
   963
        dist_eqs = dist_eqs,
35459
3d8acfae6fb8 move proofs of when_rews intro domain_constructors.ML
huffman
parents: 35458
diff changeset
   964
        cases = cases,
35460
8cb42aa19358 move definition of discriminators to domain_constructors.ML
huffman
parents: 35459
diff changeset
   965
        sel_rews = sel_thms,
35462
f5461b02d754 move definition of match combinators to domain_constructors.ML
huffman
parents: 35461
diff changeset
   966
        dis_rews = dis_thms,
40017
575d3aa1f3c5 include iso_info as part of constr_info type
huffman
parents: 40016
diff changeset
   967
        match_rews = match_thms
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   968
      }
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
   969
  in
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
   970
    (result, thy)
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   971
  end
35444
73f645fdd4ff reorganizing domain package code (in progress)
huffman
parents:
diff changeset
   972
40832
4352ca878c41 remove gratuitous semicolons from ML code
huffman
parents: 40774
diff changeset
   973
end