src/HOLCF/Tools/Domain/domain_take_proofs.ML
author huffman
Mon, 08 Mar 2010 09:37:03 -0800
changeset 35654 7a15e181bf3b
parent 35651 5dd352a85464
child 35655 e8e4af6da819
permissions -rw-r--r--
move proofs of reach and take lemmas to domain_take_proofs.ML
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
     1
(*  Title:      HOLCF/Tools/domain/domain_take_proofs.ML
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
     2
    Author:     Brian Huffman
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
     3
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
     4
Defines take functions for the given domain equation
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
     5
and proves related theorems.
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
     6
*)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
     7
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
     8
signature DOMAIN_TAKE_PROOFS =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
     9
sig
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    10
  type iso_info =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    11
    {
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    12
      absT : typ,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    13
      repT : typ,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    14
      abs_const : term,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    15
      rep_const : term,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    16
      abs_inverse : thm,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    17
      rep_inverse : thm
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    18
    }
35651
5dd352a85464 add type take_info
huffman
parents: 35650
diff changeset
    19
  type take_info =
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    20
    { take_consts : term list,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    21
      take_defs : thm list,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    22
      chain_take_thms : thm list,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    23
      take_0_thms : thm list,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    24
      take_Suc_thms : thm list,
35515
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
    25
      deflation_take_thms : thm list,
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
    26
      finite_consts : term list,
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
    27
      finite_defs : thm list
35651
5dd352a85464 add type take_info
huffman
parents: 35650
diff changeset
    28
    }
5dd352a85464 add type take_info
huffman
parents: 35650
diff changeset
    29
  val define_take_functions :
5dd352a85464 add type take_info
huffman
parents: 35650
diff changeset
    30
    (binding * iso_info) list -> theory -> take_info * theory
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    31
35654
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
    32
  val add_lub_take_theorems :
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
    33
    (binding * iso_info) list -> take_info -> thm list ->
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
    34
    theory -> unit * theory
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
    35
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    36
  val map_of_typ :
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    37
    theory -> (typ * term) list -> typ -> term
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    38
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    39
  val add_map_function :
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    40
    (string * string * thm) -> theory -> theory
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    41
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    42
  val get_map_tab : theory -> string Symtab.table
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    43
  val get_deflation_thms : theory -> thm list
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    44
end;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    45
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    46
structure Domain_Take_Proofs : DOMAIN_TAKE_PROOFS =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    47
struct
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    48
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    49
type iso_info =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    50
  {
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    51
    absT : typ,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    52
    repT : typ,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    53
    abs_const : term,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    54
    rep_const : term,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    55
    abs_inverse : thm,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    56
    rep_inverse : thm
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    57
  };
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    58
35651
5dd352a85464 add type take_info
huffman
parents: 35650
diff changeset
    59
type take_info =
5dd352a85464 add type take_info
huffman
parents: 35650
diff changeset
    60
  { take_consts : term list,
5dd352a85464 add type take_info
huffman
parents: 35650
diff changeset
    61
    take_defs : thm list,
5dd352a85464 add type take_info
huffman
parents: 35650
diff changeset
    62
    chain_take_thms : thm list,
5dd352a85464 add type take_info
huffman
parents: 35650
diff changeset
    63
    take_0_thms : thm list,
5dd352a85464 add type take_info
huffman
parents: 35650
diff changeset
    64
    take_Suc_thms : thm list,
5dd352a85464 add type take_info
huffman
parents: 35650
diff changeset
    65
    deflation_take_thms : thm list,
5dd352a85464 add type take_info
huffman
parents: 35650
diff changeset
    66
    finite_consts : term list,
5dd352a85464 add type take_info
huffman
parents: 35650
diff changeset
    67
    finite_defs : thm list
5dd352a85464 add type take_info
huffman
parents: 35650
diff changeset
    68
  };
5dd352a85464 add type take_info
huffman
parents: 35650
diff changeset
    69
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    70
val beta_ss =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    71
  HOL_basic_ss
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    72
    addsimps simp_thms
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    73
    addsimps [@{thm beta_cfun}]
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    74
    addsimprocs [@{simproc cont_proc}];
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    75
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    76
val beta_tac = simp_tac beta_ss;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    77
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    78
(******************************************************************************)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    79
(******************************** theory data *********************************)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    80
(******************************************************************************)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    81
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    82
structure MapData = Theory_Data
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    83
(
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    84
  (* constant names like "foo_map" *)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    85
  type T = string Symtab.table;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    86
  val empty = Symtab.empty;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    87
  val extend = I;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    88
  fun merge data = Symtab.merge (K true) data;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    89
);
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    90
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    91
structure DeflMapData = Theory_Data
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    92
(
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    93
  (* theorems like "deflation a ==> deflation (foo_map$a)" *)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    94
  type T = thm list;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    95
  val empty = [];
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    96
  val extend = I;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    97
  val merge = Thm.merge_thms;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    98
);
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
    99
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   100
fun add_map_function (tname, map_name, deflation_map_thm) =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   101
    MapData.map (Symtab.insert (K true) (tname, map_name))
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   102
    #> DeflMapData.map (Thm.add_thm deflation_map_thm);
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   103
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   104
val get_map_tab = MapData.get;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   105
val get_deflation_thms = DeflMapData.get;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   106
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   107
(******************************************************************************)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   108
(************************** building types and terms **************************)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   109
(******************************************************************************)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   110
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   111
open HOLCF_Library;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   112
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   113
infixr 6 ->>;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   114
infix -->>;
35515
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   115
infix 9 `;
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   116
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   117
fun mapT (T as Type (_, Ts)) =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   118
    (map (fn T => T ->> T) Ts) -->> (T ->> T)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   119
  | mapT T = T ->> T;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   120
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   121
fun mk_deflation t =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   122
  Const (@{const_name deflation}, Term.fastype_of t --> boolT) $ t;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   123
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   124
fun mk_eqs (t, u) = HOLogic.mk_Trueprop (HOLogic.mk_eq (t, u));
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   125
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   126
(******************************************************************************)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   127
(****************************** isomorphism info ******************************)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   128
(******************************************************************************)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   129
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   130
fun deflation_abs_rep (info : iso_info) : thm =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   131
  let
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   132
    val abs_iso = #abs_inverse info;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   133
    val rep_iso = #rep_inverse info;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   134
    val thm = @{thm deflation_abs_rep} OF [abs_iso, rep_iso];
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   135
  in
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   136
    Drule.export_without_context thm
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   137
  end
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   138
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   139
(******************************************************************************)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   140
(********************* building map functions over types **********************)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   141
(******************************************************************************)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   142
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   143
fun map_of_typ (thy : theory) (sub : (typ * term) list) (T : typ) : term =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   144
  let
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   145
    val map_tab = get_map_tab thy;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   146
    fun auto T = T ->> T;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   147
    fun map_of T =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   148
        case AList.lookup (op =) sub T of
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   149
          SOME m => (m, true) | NONE => map_of' T
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   150
    and map_of' (T as (Type (c, Ts))) =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   151
        (case Symtab.lookup map_tab c of
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   152
          SOME map_name =>
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   153
          let
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   154
            val map_type = map auto Ts -->> auto T;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   155
            val (ms, bs) = map_split map_of Ts;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   156
          in
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   157
            if exists I bs
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   158
            then (list_ccomb (Const (map_name, map_type), ms), true)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   159
            else (mk_ID T, false)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   160
          end
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   161
        | NONE => (mk_ID T, false))
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   162
      | map_of' T = (mk_ID T, false);
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   163
  in
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   164
    fst (map_of T)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   165
  end;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   166
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   167
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   168
(******************************************************************************)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   169
(********************* declaring definitions and theorems *********************)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   170
(******************************************************************************)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   171
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   172
fun define_const
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   173
    (bind : binding, rhs : term)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   174
    (thy : theory)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   175
    : (term * thm) * theory =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   176
  let
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   177
    val typ = Term.fastype_of rhs;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   178
    val (const, thy) = Sign.declare_const ((bind, typ), NoSyn) thy;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   179
    val eqn = Logic.mk_equals (const, rhs);
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   180
    val def = Thm.no_attributes (Binding.suffix_name "_def" bind, eqn);
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   181
    val (def_thm, thy) = yield_singleton (PureThy.add_defs false) def thy;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   182
  in
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   183
    ((const, def_thm), thy)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   184
  end;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   185
35650
64fff18d7f08 add function add_qualified_def
huffman
parents: 35594
diff changeset
   186
fun add_qualified_def name (path, eqn) thy =
64fff18d7f08 add function add_qualified_def
huffman
parents: 35594
diff changeset
   187
    thy
64fff18d7f08 add function add_qualified_def
huffman
parents: 35594
diff changeset
   188
    |> Sign.add_path path
64fff18d7f08 add function add_qualified_def
huffman
parents: 35594
diff changeset
   189
    |> yield_singleton (PureThy.add_defs false)
64fff18d7f08 add function add_qualified_def
huffman
parents: 35594
diff changeset
   190
        (Thm.no_attributes (Binding.name name, eqn))
64fff18d7f08 add function add_qualified_def
huffman
parents: 35594
diff changeset
   191
    ||> Sign.parent_path;
64fff18d7f08 add function add_qualified_def
huffman
parents: 35594
diff changeset
   192
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   193
fun add_qualified_thm name (path, thm) thy =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   194
    thy
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   195
    |> Sign.add_path path
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   196
    |> yield_singleton PureThy.add_thms
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   197
        (Thm.no_attributes (Binding.name name, thm))
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   198
    ||> Sign.parent_path;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   199
35573
bd8b50e76e94 add function add_qualified_simp_thm
huffman
parents: 35572
diff changeset
   200
fun add_qualified_simp_thm name (path, thm) thy =
bd8b50e76e94 add function add_qualified_simp_thm
huffman
parents: 35572
diff changeset
   201
    thy
bd8b50e76e94 add function add_qualified_simp_thm
huffman
parents: 35572
diff changeset
   202
    |> Sign.add_path path
bd8b50e76e94 add function add_qualified_simp_thm
huffman
parents: 35572
diff changeset
   203
    |> yield_singleton PureThy.add_thms
bd8b50e76e94 add function add_qualified_simp_thm
huffman
parents: 35572
diff changeset
   204
        ((Binding.name name, thm), [Simplifier.simp_add])
bd8b50e76e94 add function add_qualified_simp_thm
huffman
parents: 35572
diff changeset
   205
    ||> Sign.parent_path;
bd8b50e76e94 add function add_qualified_simp_thm
huffman
parents: 35572
diff changeset
   206
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   207
(******************************************************************************)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   208
(************************** defining take functions ***************************)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   209
(******************************************************************************)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   210
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   211
fun define_take_functions
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   212
    (spec : (binding * iso_info) list)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   213
    (thy : theory) =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   214
  let
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   215
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   216
    (* retrieve components of spec *)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   217
    val dom_binds = map fst spec;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   218
    val iso_infos = map snd spec;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   219
    val dom_eqns = map (fn x => (#absT x, #repT x)) iso_infos;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   220
    val rep_abs_consts = map (fn x => (#rep_const x, #abs_const x)) iso_infos;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   221
    val dnames = map Binding.name_of dom_binds;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   222
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   223
    (* get table of map functions *)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   224
    val map_tab = MapData.get thy;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   225
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   226
    fun mk_projs []      t = []
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   227
      | mk_projs (x::[]) t = [(x, t)]
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   228
      | mk_projs (x::xs) t = (x, mk_fst t) :: mk_projs xs (mk_snd t);
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   229
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   230
    fun mk_cfcomp2 ((rep_const, abs_const), f) =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   231
        mk_cfcomp (abs_const, mk_cfcomp (f, rep_const));
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   232
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   233
    (* define take functional *)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   234
    val newTs : typ list = map fst dom_eqns;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   235
    val copy_arg_type = mk_tupleT (map (fn T => T ->> T) newTs);
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   236
    val copy_arg = Free ("f", copy_arg_type);
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   237
    val copy_args = map snd (mk_projs dom_binds copy_arg);
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   238
    fun one_copy_rhs (rep_abs, (lhsT, rhsT)) =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   239
      let
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   240
        val body = map_of_typ thy (newTs ~~ copy_args) rhsT;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   241
      in
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   242
        mk_cfcomp2 (rep_abs, body)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   243
      end;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   244
    val take_functional =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   245
        big_lambda copy_arg
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   246
          (mk_tuple (map one_copy_rhs (rep_abs_consts ~~ dom_eqns)));
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   247
    val take_rhss =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   248
      let
35557
5da670d57118 uniformly use variable names m and n in take-related lemmas; use export_without_context where appropriate
huffman
parents: 35555
diff changeset
   249
        val n = Free ("n", HOLogic.natT);
5da670d57118 uniformly use variable names m and n in take-related lemmas; use export_without_context where appropriate
huffman
parents: 35555
diff changeset
   250
        val rhs = mk_iterate (n, take_functional);
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   251
      in
35557
5da670d57118 uniformly use variable names m and n in take-related lemmas; use export_without_context where appropriate
huffman
parents: 35555
diff changeset
   252
        map (lambda n o snd) (mk_projs dom_binds rhs)
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   253
      end;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   254
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   255
    (* define take constants *)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   256
    fun define_take_const ((tbind, take_rhs), (lhsT, rhsT)) thy =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   257
      let
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   258
        val take_type = HOLogic.natT --> lhsT ->> lhsT;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   259
        val take_bind = Binding.suffix_name "_take" tbind;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   260
        val (take_const, thy) =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   261
          Sign.declare_const ((take_bind, take_type), NoSyn) thy;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   262
        val take_eqn = Logic.mk_equals (take_const, take_rhs);
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   263
        val (take_def_thm, thy) =
35650
64fff18d7f08 add function add_qualified_def
huffman
parents: 35594
diff changeset
   264
            add_qualified_def "take_def"
64fff18d7f08 add function add_qualified_def
huffman
parents: 35594
diff changeset
   265
             (Binding.name_of tbind, take_eqn) thy;
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   266
      in ((take_const, take_def_thm), thy) end;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   267
    val ((take_consts, take_defs), thy) = thy
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   268
      |> fold_map define_take_const (dom_binds ~~ take_rhss ~~ dom_eqns)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   269
      |>> ListPair.unzip;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   270
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   271
    (* prove chain_take lemmas *)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   272
    fun prove_chain_take (take_const, dname) thy =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   273
      let
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   274
        val goal = mk_trp (mk_chain take_const);
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   275
        val rules = take_defs @ @{thms chain_iterate ch2ch_fst ch2ch_snd};
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   276
        val tac = simp_tac (HOL_basic_ss addsimps rules) 1;
35572
44eeda8cb708 generate lemma take_below, declare chain_take [simp]
huffman
parents: 35557
diff changeset
   277
        val thm = Goal.prove_global thy [] [] goal (K tac);
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   278
      in
35573
bd8b50e76e94 add function add_qualified_simp_thm
huffman
parents: 35572
diff changeset
   279
        add_qualified_simp_thm "chain_take" (dname, thm) thy
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   280
      end;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   281
    val (chain_take_thms, thy) =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   282
      fold_map prove_chain_take (take_consts ~~ dnames) thy;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   283
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   284
    (* prove take_0 lemmas *)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   285
    fun prove_take_0 ((take_const, dname), (lhsT, rhsT)) thy =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   286
      let
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   287
        val lhs = take_const $ @{term "0::nat"};
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   288
        val goal = mk_eqs (lhs, mk_bottom (lhsT ->> lhsT));
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   289
        val rules = take_defs @ @{thms iterate_0 fst_strict snd_strict};
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   290
        val tac = simp_tac (HOL_basic_ss addsimps rules) 1;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   291
        val take_0_thm = Goal.prove_global thy [] [] goal (K tac);
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   292
      in
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   293
        add_qualified_thm "take_0" (dname, take_0_thm) thy
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   294
      end;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   295
    val (take_0_thms, thy) =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   296
      fold_map prove_take_0 (take_consts ~~ dnames ~~ dom_eqns) thy;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   297
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   298
    (* prove take_Suc lemmas *)
35557
5da670d57118 uniformly use variable names m and n in take-related lemmas; use export_without_context where appropriate
huffman
parents: 35555
diff changeset
   299
    val n = Free ("n", natT);
5da670d57118 uniformly use variable names m and n in take-related lemmas; use export_without_context where appropriate
huffman
parents: 35555
diff changeset
   300
    val take_is = map (fn t => t $ n) take_consts;
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   301
    fun prove_take_Suc
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   302
          (((take_const, rep_abs), dname), (lhsT, rhsT)) thy =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   303
      let
35557
5da670d57118 uniformly use variable names m and n in take-related lemmas; use export_without_context where appropriate
huffman
parents: 35555
diff changeset
   304
        val lhs = take_const $ (@{term Suc} $ n);
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   305
        val body = map_of_typ thy (newTs ~~ take_is) rhsT;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   306
        val rhs = mk_cfcomp2 (rep_abs, body);
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   307
        val goal = mk_eqs (lhs, rhs);
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   308
        val simps = @{thms iterate_Suc fst_conv snd_conv}
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   309
        val rules = take_defs @ simps;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   310
        val tac = simp_tac (beta_ss addsimps rules) 1;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   311
        val take_Suc_thm = Goal.prove_global thy [] [] goal (K tac);
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   312
      in
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   313
        add_qualified_thm "take_Suc" (dname, take_Suc_thm) thy
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   314
      end;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   315
    val (take_Suc_thms, thy) =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   316
      fold_map prove_take_Suc
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   317
        (take_consts ~~ rep_abs_consts ~~ dnames ~~ dom_eqns) thy;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   318
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   319
    (* prove deflation theorems for take functions *)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   320
    val deflation_abs_rep_thms = map deflation_abs_rep iso_infos;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   321
    val deflation_take_thm =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   322
      let
35557
5da670d57118 uniformly use variable names m and n in take-related lemmas; use export_without_context where appropriate
huffman
parents: 35555
diff changeset
   323
        val n = Free ("n", natT);
5da670d57118 uniformly use variable names m and n in take-related lemmas; use export_without_context where appropriate
huffman
parents: 35555
diff changeset
   324
        fun mk_goal take_const = mk_deflation (take_const $ n);
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   325
        val goal = mk_trp (foldr1 mk_conj (map mk_goal take_consts));
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   326
        val adm_rules =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   327
          @{thms adm_conj adm_subst [OF _ adm_deflation]
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   328
                 cont2cont_fst cont2cont_snd cont_id};
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   329
        val bottom_rules =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   330
          take_0_thms @ @{thms deflation_UU simp_thms};
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   331
        val deflation_rules =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   332
          @{thms conjI deflation_ID}
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   333
          @ deflation_abs_rep_thms
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   334
          @ DeflMapData.get thy;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   335
      in
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   336
        Goal.prove_global thy [] [] goal (fn _ =>
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   337
         EVERY
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   338
          [rtac @{thm nat.induct} 1,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   339
           simp_tac (HOL_basic_ss addsimps bottom_rules) 1,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   340
           asm_simp_tac (HOL_basic_ss addsimps take_Suc_thms) 1,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   341
           REPEAT (etac @{thm conjE} 1
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   342
                   ORELSE resolve_tac deflation_rules 1
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   343
                   ORELSE atac 1)])
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   344
      end;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   345
    fun conjuncts [] thm = []
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   346
      | conjuncts (n::[]) thm = [(n, thm)]
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   347
      | conjuncts (n::ns) thm = let
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   348
          val thmL = thm RS @{thm conjunct1};
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   349
          val thmR = thm RS @{thm conjunct2};
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   350
        in (n, thmL):: conjuncts ns thmR end;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   351
    val (deflation_take_thms, thy) =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   352
      fold_map (add_qualified_thm "deflation_take")
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   353
        (map (apsnd Drule.export_without_context)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   354
          (conjuncts dnames deflation_take_thm)) thy;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   355
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   356
    (* prove strictness of take functions *)
35572
44eeda8cb708 generate lemma take_below, declare chain_take [simp]
huffman
parents: 35557
diff changeset
   357
    fun prove_take_strict (deflation_take, dname) thy =
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   358
      let
35572
44eeda8cb708 generate lemma take_below, declare chain_take [simp]
huffman
parents: 35557
diff changeset
   359
        val take_strict_thm =
44eeda8cb708 generate lemma take_below, declare chain_take [simp]
huffman
parents: 35557
diff changeset
   360
            Drule.export_without_context
44eeda8cb708 generate lemma take_below, declare chain_take [simp]
huffman
parents: 35557
diff changeset
   361
            (@{thm deflation_strict} OF [deflation_take]);
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   362
      in
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   363
        add_qualified_thm "take_strict" (dname, take_strict_thm) thy
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   364
      end;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   365
    val (take_strict_thms, thy) =
35572
44eeda8cb708 generate lemma take_below, declare chain_take [simp]
huffman
parents: 35557
diff changeset
   366
      fold_map prove_take_strict
44eeda8cb708 generate lemma take_below, declare chain_take [simp]
huffman
parents: 35557
diff changeset
   367
        (deflation_take_thms ~~ dnames) thy;
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   368
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   369
    (* prove take/take rules *)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   370
    fun prove_take_take ((chain_take, deflation_take), dname) thy =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   371
      let
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   372
        val take_take_thm =
35557
5da670d57118 uniformly use variable names m and n in take-related lemmas; use export_without_context where appropriate
huffman
parents: 35555
diff changeset
   373
            Drule.export_without_context
5da670d57118 uniformly use variable names m and n in take-related lemmas; use export_without_context where appropriate
huffman
parents: 35555
diff changeset
   374
            (@{thm deflation_chain_min} OF [chain_take, deflation_take]);
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   375
      in
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   376
        add_qualified_thm "take_take" (dname, take_take_thm) thy
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   377
      end;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   378
    val (take_take_thms, thy) =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   379
      fold_map prove_take_take
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   380
        (chain_take_thms ~~ deflation_take_thms ~~ dnames) thy;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   381
35572
44eeda8cb708 generate lemma take_below, declare chain_take [simp]
huffman
parents: 35557
diff changeset
   382
    (* prove take_below rules *)
44eeda8cb708 generate lemma take_below, declare chain_take [simp]
huffman
parents: 35557
diff changeset
   383
    fun prove_take_below (deflation_take, dname) thy =
44eeda8cb708 generate lemma take_below, declare chain_take [simp]
huffman
parents: 35557
diff changeset
   384
      let
44eeda8cb708 generate lemma take_below, declare chain_take [simp]
huffman
parents: 35557
diff changeset
   385
        val take_below_thm =
44eeda8cb708 generate lemma take_below, declare chain_take [simp]
huffman
parents: 35557
diff changeset
   386
            Drule.export_without_context
44eeda8cb708 generate lemma take_below, declare chain_take [simp]
huffman
parents: 35557
diff changeset
   387
            (@{thm deflation.below} OF [deflation_take]);
44eeda8cb708 generate lemma take_below, declare chain_take [simp]
huffman
parents: 35557
diff changeset
   388
      in
44eeda8cb708 generate lemma take_below, declare chain_take [simp]
huffman
parents: 35557
diff changeset
   389
        add_qualified_thm "take_below" (dname, take_below_thm) thy
44eeda8cb708 generate lemma take_below, declare chain_take [simp]
huffman
parents: 35557
diff changeset
   390
      end;
44eeda8cb708 generate lemma take_below, declare chain_take [simp]
huffman
parents: 35557
diff changeset
   391
    val (take_below_thms, thy) =
44eeda8cb708 generate lemma take_below, declare chain_take [simp]
huffman
parents: 35557
diff changeset
   392
      fold_map prove_take_below
44eeda8cb708 generate lemma take_below, declare chain_take [simp]
huffman
parents: 35557
diff changeset
   393
        (deflation_take_thms ~~ dnames) thy;
44eeda8cb708 generate lemma take_below, declare chain_take [simp]
huffman
parents: 35557
diff changeset
   394
35515
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   395
    (* define finiteness predicates *)
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   396
    fun define_finite_const ((tbind, take_const), (lhsT, rhsT)) thy =
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   397
      let
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   398
        val finite_type = lhsT --> boolT;
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   399
        val finite_bind = Binding.suffix_name "_finite" tbind;
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   400
        val (finite_const, thy) =
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   401
          Sign.declare_const ((finite_bind, finite_type), NoSyn) thy;
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   402
        val x = Free ("x", lhsT);
35557
5da670d57118 uniformly use variable names m and n in take-related lemmas; use export_without_context where appropriate
huffman
parents: 35555
diff changeset
   403
        val n = Free ("n", natT);
35515
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   404
        val finite_rhs =
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   405
          lambda x (HOLogic.exists_const natT $
35557
5da670d57118 uniformly use variable names m and n in take-related lemmas; use export_without_context where appropriate
huffman
parents: 35555
diff changeset
   406
            (lambda n (mk_eq (mk_capply (take_const $ n, x), x))));
35515
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   407
        val finite_eqn = Logic.mk_equals (finite_const, finite_rhs);
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   408
        val (finite_def_thm, thy) =
35650
64fff18d7f08 add function add_qualified_def
huffman
parents: 35594
diff changeset
   409
            add_qualified_def "finite_def"
64fff18d7f08 add function add_qualified_def
huffman
parents: 35594
diff changeset
   410
             (Binding.name_of tbind, finite_eqn) thy;
35515
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   411
      in ((finite_const, finite_def_thm), thy) end;
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   412
    val ((finite_consts, finite_defs), thy) = thy
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   413
      |> fold_map define_finite_const (dom_binds ~~ take_consts ~~ dom_eqns)
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   414
      |>> ListPair.unzip;
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   415
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   416
    val result =
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   417
      {
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   418
        take_consts = take_consts,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   419
        take_defs = take_defs,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   420
        chain_take_thms = chain_take_thms,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   421
        take_0_thms = take_0_thms,
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   422
        take_Suc_thms = take_Suc_thms,
35515
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   423
        deflation_take_thms = deflation_take_thms,
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   424
        finite_consts = finite_consts,
d631dc53ede0 move definition of finiteness predicate into domain_take_proofs.ML
huffman
parents: 35514
diff changeset
   425
        finite_defs = finite_defs
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   426
      };
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   427
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   428
  in
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   429
    (result, thy)
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   430
  end;
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   431
35654
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   432
fun add_lub_take_theorems
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   433
    (spec : (binding * iso_info) list)
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   434
    (take_info : take_info)
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   435
    (lub_take_thms : thm list)
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   436
    (thy : theory) =
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   437
  let
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   438
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   439
    (* retrieve components of spec *)
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   440
    val dom_binds = map fst spec;
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   441
    val iso_infos = map snd spec;
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   442
    val dom_eqns = map (fn x => (#absT x, #repT x)) iso_infos;
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   443
    val dnames = map Binding.name_of dom_binds;
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   444
    val {chain_take_thms, deflation_take_thms, ...} = take_info;
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   445
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   446
    (* prove take lemmas *)
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   447
    fun prove_take_lemma ((chain_take, lub_take), dname) thy =
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   448
      let
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   449
        val take_lemma =
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   450
            Drule.export_without_context
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   451
              (@{thm lub_ID_take_lemma} OF [chain_take, lub_take]);
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   452
      in
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   453
        add_qualified_thm "take_lemma" (dname, take_lemma) thy
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   454
      end;
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   455
    val (take_lemma_thms, thy) =
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   456
      fold_map prove_take_lemma
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   457
        (chain_take_thms ~~ lub_take_thms ~~ dnames) thy;
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   458
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   459
    (* prove reach lemmas *)
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   460
    fun prove_reach_lemma ((chain_take, lub_take), dname) thy =
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   461
      let
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   462
        val thm =
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   463
            Drule.export_without_context
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   464
              (@{thm lub_ID_reach} OF [chain_take, lub_take]);
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   465
      in
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   466
        add_qualified_thm "reach" (dname, thm) thy
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   467
      end;
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   468
    val (reach_thms, thy) =
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   469
      fold_map prove_reach_lemma
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   470
        (chain_take_thms ~~ lub_take_thms ~~ dnames) thy;
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   471
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   472
    val result = ();
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   473
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   474
  in
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   475
    (result, thy)
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   476
  end;
7a15e181bf3b move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents: 35651
diff changeset
   477
35514
a2cfa413eaab move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
diff changeset
   478
end;