src/HOL/Tools/BNF/bnf_gfp_util.ML
author haftmann
Fri, 27 Jun 2025 08:09:26 +0200
changeset 82775 61c39a9e5415
parent 75624 22d1c5f2b9f4
permissions -rw-r--r--
typo
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55061
a0adf838e2d1 adjusted comments
blanchet
parents: 55060
diff changeset
     1
(*  Title:      HOL/Tools/BNF/bnf_gfp_util.ML
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     2
    Author:     Dmitriy Traytel, TU Muenchen
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 69593
diff changeset
     3
    Author:     Jan van Brügge, TU Muenchen
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 69593
diff changeset
     4
    Copyright   2012, 2022
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     5
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     6
Library for the codatatype construction.
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     7
*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     8
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     9
signature BNF_GFP_UTIL =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    10
sig
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    11
  val mk_rec_simps: int -> thm -> thm list -> thm list list
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    12
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    13
  val dest_listT: typ -> typ
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    14
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    15
  val mk_Cons: term -> term -> term
55966
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55578
diff changeset
    16
  val mk_InN: typ list -> term -> int -> term
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    17
  val mk_Shift: term -> term -> term
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    18
  val mk_Succ: term -> term -> term
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    19
  val mk_Times: term * term -> term
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    20
  val mk_append: term * term -> term
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    21
  val mk_congruent: term -> term -> term
51447
a19e973fa2cf eliminate duplicated constant (diag vs. Id_on)
traytel
parents: 50058
diff changeset
    22
  val mk_Id_on: term -> term
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51447
diff changeset
    23
  val mk_in_rel: term -> term
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    24
  val mk_equiv: term -> term -> term
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    25
  val mk_fromCard: term -> term -> term
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    26
  val mk_proj: term -> term
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    27
  val mk_quotient: term -> term -> term
55413
a8e96847523c adapted theories to '{case,rec}_{list,option}' names
blanchet
parents: 55061
diff changeset
    28
  val mk_rec_list: term -> term -> term
55415
05f5fdb8d093 renamed 'nat_{case,rec}' to '{case,rec}_nat'
blanchet
parents: 55413
diff changeset
    29
  val mk_rec_nat: term -> term -> term
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    30
  val mk_shift: term -> term -> term
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    31
  val mk_size: term -> term
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    32
  val mk_toCard: term -> term -> term
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    33
  val mk_undefined: typ -> term
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    34
  val mk_univ: term -> term
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 69593
diff changeset
    35
  val mk_card_suc: term -> term
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    36
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    37
  val mk_specN: int -> thm -> thm
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    38
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    39
  val mk_InN_Field: int -> int -> thm
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    40
  val mk_InN_inject: int -> int -> thm
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    41
  val mk_InN_not_InM: int -> int -> thm
55966
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55578
diff changeset
    42
  val mk_sumEN: int -> thm
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    43
end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    44
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    45
structure BNF_GFP_Util : BNF_GFP_UTIL =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    46
struct
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    47
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    48
open BNF_Util
55966
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55578
diff changeset
    49
open BNF_FP_Util
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    50
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 55966
diff changeset
    51
val mk_append = HOLogic.mk_binop \<^const_name>\<open>append\<close>;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    52
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    53
fun mk_equiv B R =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 55966
diff changeset
    54
  Const (\<^const_name>\<open>equiv\<close>, fastype_of B --> fastype_of R --> HOLogic.boolT) $ B $ R;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    55
55966
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55578
diff changeset
    56
fun mk_InN [_] t 1 = t
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55578
diff changeset
    57
  | mk_InN (_ :: Ts) t 1 = mk_Inl (mk_sumTN Ts) t
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55578
diff changeset
    58
  | mk_InN (LT :: Ts) t m = mk_Inr LT (mk_InN Ts t (m - 1))
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55578
diff changeset
    59
  | mk_InN Ts t _ = raise TYPE ("mk_InN", Ts, [t]);
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55578
diff changeset
    60
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    61
fun mk_Sigma (A, B) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    62
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    63
    val AT = fastype_of A;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    64
    val BT = fastype_of B;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    65
    val ABT = mk_relT (HOLogic.dest_setT AT, HOLogic.dest_setT (range_type BT));
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 55966
diff changeset
    66
  in Const (\<^const_name>\<open>Sigma\<close>, AT --> BT --> ABT) $ A $ B end;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    67
51447
a19e973fa2cf eliminate duplicated constant (diag vs. Id_on)
traytel
parents: 50058
diff changeset
    68
fun mk_Id_on A =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    69
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    70
    val AT = fastype_of A;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    71
    val AAT = mk_relT (HOLogic.dest_setT AT, HOLogic.dest_setT AT);
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 55966
diff changeset
    72
  in Const (\<^const_name>\<open>Id_on\<close>, AT --> AAT) $ A end;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    73
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51447
diff changeset
    74
fun mk_in_rel R =
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51447
diff changeset
    75
  let
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51447
diff changeset
    76
    val ((A, B), RT) = `dest_relT (fastype_of R);
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 55966
diff changeset
    77
  in Const (\<^const_name>\<open>in_rel\<close>, RT --> mk_pred2T A B) $ R end;
51893
596baae88a88 got rid of the set based relator---use (binary) predicate based relator instead
traytel
parents: 51447
diff changeset
    78
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    79
fun mk_Times (A, B) =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    80
  let val AT = HOLogic.dest_setT (fastype_of A);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    81
  in mk_Sigma (A, Term.absdummy AT B) end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    82
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 55966
diff changeset
    83
fun dest_listT (Type (\<^type_name>\<open>list\<close>, [T])) = T
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    84
  | dest_listT T = raise TYPE ("dest_setT: set type expected", [T], []);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    85
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    86
fun mk_Succ Kl kl =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    87
  let val T = fastype_of kl;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    88
  in
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 55966
diff changeset
    89
    Const (\<^const_name>\<open>Succ\<close>,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    90
      HOLogic.mk_setT T --> T --> HOLogic.mk_setT (dest_listT T)) $ Kl $ kl
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    91
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    92
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    93
fun mk_Shift Kl k =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    94
  let val T = fastype_of Kl;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    95
  in
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 55966
diff changeset
    96
    Const (\<^const_name>\<open>Shift\<close>, T --> dest_listT (HOLogic.dest_setT T) --> T) $ Kl $ k
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    97
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    98
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    99
fun mk_shift lab k =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   100
  let val T = fastype_of lab;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   101
  in
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 55966
diff changeset
   102
    Const (\<^const_name>\<open>shift\<close>, T --> dest_listT (Term.domain_type T) --> T) $ lab $ k
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   103
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   104
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   105
fun mk_toCard A r =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   106
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   107
    val AT = fastype_of A;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   108
    val rT = fastype_of r;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   109
  in
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 55966
diff changeset
   110
    Const (\<^const_name>\<open>toCard\<close>,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   111
      AT --> rT --> HOLogic.dest_setT AT --> fst (dest_relT rT)) $ A $ r
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   112
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   113
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   114
fun mk_fromCard A r =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   115
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   116
    val AT = fastype_of A;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   117
    val rT = fastype_of r;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   118
  in
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 55966
diff changeset
   119
    Const (\<^const_name>\<open>fromCard\<close>,
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   120
      AT --> rT --> fst (dest_relT rT) --> HOLogic.dest_setT AT) $ A $ r
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   121
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   122
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   123
fun mk_Cons x xs =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   124
  let val T = fastype_of xs;
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 55966
diff changeset
   125
  in Const (\<^const_name>\<open>Cons\<close>, dest_listT T --> T --> T) $ x $ xs end;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   126
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   127
fun mk_size t = HOLogic.size_const (fastype_of t) $ t;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   128
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   129
fun mk_quotient A R =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   130
  let val T = fastype_of A;
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 55966
diff changeset
   131
  in Const (\<^const_name>\<open>quotient\<close>, T --> fastype_of R --> HOLogic.mk_setT T) $ A $ R end;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   132
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   133
fun mk_proj R =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   134
  let val ((AT, BT), T) = `dest_relT (fastype_of R);
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 55966
diff changeset
   135
  in Const (\<^const_name>\<open>proj\<close>, T --> AT --> HOLogic.mk_setT BT) $ R end;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   136
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   137
fun mk_univ f =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   138
  let val ((AT, BT), T) = `dest_funT (fastype_of f);
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 55966
diff changeset
   139
  in Const (\<^const_name>\<open>univ\<close>, T --> HOLogic.mk_setT AT --> BT) $ f end;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   140
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   141
fun mk_congruent R f =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 55966
diff changeset
   142
  Const (\<^const_name>\<open>congruent\<close>, fastype_of R --> fastype_of f --> HOLogic.boolT) $ R $ f;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   143
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 55966
diff changeset
   144
fun mk_undefined T = Const (\<^const_name>\<open>undefined\<close>, T);
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   145
75624
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 69593
diff changeset
   146
fun mk_card_suc t =
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 69593
diff changeset
   147
  let
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 69593
diff changeset
   148
    val T = fst (dest_relT (fastype_of t))
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 69593
diff changeset
   149
    val T' = Type (\<^type_name>\<open>suc\<close>, [T])
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 69593
diff changeset
   150
  in Const (\<^const_name>\<open>card_suc\<close>, mk_relT (T, T) --> mk_relT (T', T')) $ t end;
22d1c5f2b9f4 strict bounds for BNFs (by Jan van Brügge)
traytel
parents: 69593
diff changeset
   151
55415
05f5fdb8d093 renamed 'nat_{case,rec}' to '{case,rec}_nat'
blanchet
parents: 55413
diff changeset
   152
fun mk_rec_nat Zero Suc =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   153
  let val T = fastype_of Zero;
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 55966
diff changeset
   154
  in Const (\<^const_name>\<open>old.rec_nat\<close>, T --> fastype_of Suc --> HOLogic.natT --> T) $ Zero $ Suc end;
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   155
55413
a8e96847523c adapted theories to '{case,rec}_{list,option}' names
blanchet
parents: 55061
diff changeset
   156
fun mk_rec_list Nil Cons =
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   157
  let
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   158
    val T = fastype_of Nil;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   159
    val (U, consT) = `(Term.domain_type) (fastype_of Cons);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   160
  in
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 55966
diff changeset
   161
    Const (\<^const_name>\<open>rec_list\<close>, T --> consT --> HOLogic.listT U --> T) $ Nil $ Cons
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   162
  end;
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   163
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   164
fun mk_InN_not_InM 1 _ = @{thm Inl_not_Inr}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   165
  | mk_InN_not_InM n m =
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   166
    if n > m then mk_InN_not_InM m n RS @{thm not_sym}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   167
    else mk_InN_not_InM (n - 1) (m - 1) RS @{thm not_arg_cong_Inr};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   168
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   169
fun mk_InN_Field 1 1 = @{thm TrueE[OF TrueI]}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   170
  | mk_InN_Field _ 1 = @{thm Inl_Field_csum}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   171
  | mk_InN_Field 2 2 = @{thm Inr_Field_csum}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   172
  | mk_InN_Field n m = mk_InN_Field (n - 1) (m - 1) RS @{thm Inr_Field_csum};
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   173
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   174
fun mk_InN_inject 1 _ = @{thm TrueE[OF TrueI]}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   175
  | mk_InN_inject _ 1 = @{thm Inl_inject}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   176
  | mk_InN_inject 2 2 = @{thm Inr_inject}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   177
  | mk_InN_inject n m = @{thm Inr_inject} RS mk_InN_inject (n - 1) (m - 1);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   178
55966
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55578
diff changeset
   179
fun mk_sumEN 1 = @{thm one_pointE}
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55578
diff changeset
   180
  | mk_sumEN 2 = @{thm sumE}
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55578
diff changeset
   181
  | mk_sumEN n =
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55578
diff changeset
   182
    (fold (fn i => fn thm => @{thm obj_sumE_f} RSN (i, thm)) (2 upto n - 1) @{thm obj_sumE}) OF
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55578
diff changeset
   183
      replicate n (impI RS allI);
972f0aa7091b balance tuples that represent curried functions
blanchet
parents: 55578
diff changeset
   184
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   185
fun mk_specN 0 thm = thm
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   186
  | mk_specN n thm = mk_specN (n - 1) (thm RS spec);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   187
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   188
fun mk_rec_simps n rec_thm defs = map (fn i =>
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   189
  map (fn def => def RS rec_thm RS mk_nthI n i RS fun_cong) defs) (1 upto n);
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   190
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   191
end;