| author | paulson <lp15@cam.ac.uk> | 
| Mon, 27 Feb 2023 17:09:59 +0000 | |
| changeset 77406 | c2013f617a70 | 
| parent 76743 | d33fc5228aae | 
| permissions | -rw-r--r-- | 
| 55061 | 1 | (* Title: HOL/Tools/BNF/bnf_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 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 3 | Copyright 2012 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 4 | |
| 49282 | 5 | Library for bounded natural functors. | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 6 | *) | 
| 
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 | signature BNF_UTIL = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 9 | sig | 
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: 
53888diff
changeset | 10 | include CTR_SUGAR_UTIL | 
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 11 | include BNF_FP_REC_SUGAR_UTIL | 
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: 
53888diff
changeset | 12 | |
| 64413 | 13 | val transfer_plugin: string | 
| 14 | ||
| 63797 | 15 | val unflatt: 'a list list list -> 'b list -> 'b list list list | 
| 16 | val unflattt: 'a list list list list -> 'b list -> 'b list list list list | |
| 17 | ||
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 18 | val mk_TFreess: int list -> Proof.context -> typ list list * Proof.context | 
| 49177 
db8ce685073f
introduced and used "mk_Freesss", and simplified "mk_Freess(')"
 blanchet parents: 
49176diff
changeset | 19 | val mk_Freesss: string -> typ list list list -> Proof.context -> | 
| 
db8ce685073f
introduced and used "mk_Freesss", and simplified "mk_Freess(')"
 blanchet parents: 
49176diff
changeset | 20 | term list list list * Proof.context | 
| 49200 | 21 | val mk_Freessss: string -> typ list list list list -> Proof.context -> | 
| 22 | term list list list list * Proof.context | |
| 53034 
6067703399ad
moved library function where it belongs, and used Dmitriy's inside-out implementation
 blanchet parents: 
52985diff
changeset | 23 | val nonzero_string_of_int: int -> string | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 24 | |
| 53036 | 25 | val binder_fun_types: typ -> typ list | 
| 26 | val body_fun_type: typ -> typ | |
| 53037 | 27 | val strip_fun_type: typ -> typ list * typ | 
| 49395 
323414474c1f
use strip_typeN in bnf_def (instead of repairing strip_type)
 traytel parents: 
49366diff
changeset | 28 | val strip_typeN: int -> typ -> typ list * typ | 
| 
323414474c1f
use strip_typeN in bnf_def (instead of repairing strip_type)
 traytel parents: 
49366diff
changeset | 29 | |
| 49463 | 30 | val mk_pred2T: typ -> typ -> typ | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 31 | val mk_relT: typ * typ -> typ | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 32 | val dest_relT: typ -> typ * typ | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 33 | val dest_pred2T: typ -> typ * typ | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 34 | val mk_sumT: typ * typ -> typ | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 35 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 36 | val ctwo: term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 37 | val fst_const: typ -> term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 38 | val snd_const: typ -> term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 39 | val Id_const: typ -> term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 40 | |
| 63824 | 41 | val enforce_type: Proof.context -> (typ -> typ) -> typ -> term -> term | 
| 42 | ||
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 43 | val mk_Ball: term -> term -> term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 44 | val mk_Bex: term -> term -> term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 45 | val mk_Card_order: term -> term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 46 | val mk_Field: term -> term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 47 | val mk_Gr: term -> term -> term | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 48 | val mk_Grp: term -> term -> term | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 49 | val mk_UNION: term -> term -> term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 50 | val mk_Union: typ -> term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 51 | val mk_card_binop: string -> (typ * typ -> typ) -> term -> term -> term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 52 | val mk_card_of: term -> term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 53 | val mk_card_order: term -> term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 54 | val mk_cexp: term -> term -> term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 55 | val mk_cinfinite: term -> term | 
| 75624 | 56 | val mk_regularCard: term -> term | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 57 | val mk_collect: term list -> typ -> term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 58 | val mk_converse: term -> term | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 59 | val mk_conversep: term -> term | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 60 | val mk_cprod: term -> term -> term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 61 | val mk_csum: term -> term -> term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 62 | val mk_dir_image: term -> term -> term | 
| 62324 | 63 | val mk_eq_onp: term -> term | 
| 55945 | 64 | val mk_rel_fun: term -> term -> term | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 65 | val mk_image: term -> term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 66 | val mk_in: term list -> term list -> typ -> term | 
| 56635 | 67 | val mk_inj: term -> term | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 68 | val mk_leq: term -> term -> term | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 69 | val mk_ordLeq: term -> term -> term | 
| 75624 | 70 | val mk_ordLess: term -> term -> term | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 71 | val mk_rel_comp: term * term -> term | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 72 | val mk_rel_compp: term * term -> term | 
| 55803 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 traytel parents: 
55772diff
changeset | 73 | val mk_vimage2p: term -> term -> term | 
| 61240 | 74 | val mk_reflp: term -> term | 
| 75 | val mk_symp: term -> term | |
| 76 | val mk_transp: term -> term | |
| 63851 | 77 | val mk_union: term * term -> term | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 78 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 79 | (*parameterized terms*) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 80 | val mk_nthN: int -> term -> int -> term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 81 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 82 | (*parameterized thms*) | 
| 56765 | 83 | val prod_injectD: thm | 
| 84 | val prod_injectI: thm | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 85 | val ctrans: thm | 
| 49585 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 blanchet parents: 
49536diff
changeset | 86 | val id_apply: thm | 
| 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 blanchet parents: 
49536diff
changeset | 87 | val meta_mp: thm | 
| 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 blanchet parents: 
49536diff
changeset | 88 | val meta_spec: thm | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 89 | val o_apply: thm | 
| 58446 | 90 | val rel_funD: thm | 
| 91 | val rel_funI: thm | |
| 49488 
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
 traytel parents: 
49484diff
changeset | 92 | val set_mp: thm | 
| 
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
 traytel parents: 
49484diff
changeset | 93 | val set_rev_mp: thm | 
| 49490 | 94 | val subset_UNIV: thm | 
| 58446 | 95 | |
| 96 | val mk_conjIN: int -> thm | |
| 97 | val mk_nthI: int -> int -> thm | |
| 98 | val mk_nth_conv: int -> int -> thm | |
| 99 | val mk_ordLeq_csum: int -> int -> thm -> thm | |
| 75624 | 100 | val mk_ordLess_csum: int -> int -> thm -> thm | 
| 62124 | 101 | val mk_pointful: Proof.context -> thm -> thm | 
| 58446 | 102 | val mk_rel_funDN: int -> thm -> thm | 
| 103 | val mk_rel_funDN_rotated: int -> thm -> thm | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 104 | val mk_sym: thm -> thm | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 105 | val mk_trans: thm -> thm -> thm | 
| 58446 | 106 | val mk_UnIN: int -> int -> thm | 
| 107 | val mk_Un_upper: int -> int -> thm | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 108 | |
| 57205 | 109 | val is_refl_bool: term -> bool | 
| 49484 | 110 | val is_refl: thm -> bool | 
| 49605 
ea566f5e1724
avoid another brand of trivial "disc_rel" theorems (which made the simplifier loop for all single-constructor types)
 blanchet parents: 
49585diff
changeset | 111 | val is_concl_refl: thm -> bool | 
| 49484 | 112 | val no_refl: thm list -> thm list | 
| 113 | val no_reflexive: thm list -> thm list | |
| 114 | ||
| 54601 
91a1e4aa7c80
command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
 traytel parents: 
54544diff
changeset | 115 | val parse_type_args_named_constrained: (binding option * (string * string option)) list parser | 
| 62324 | 116 | val parse_map_rel_pred_bindings: (binding * binding * binding) parser | 
| 49434 
433dc7e028c8
separated registration of BNFs from bnf_def (BNFs are now stored only for bnf_def and (co)data commands)
 traytel parents: 
49425diff
changeset | 117 | |
| 49835 | 118 | val typedef: binding * (string * sort) list * mixfix -> term -> | 
| 58959 | 119 | (binding * binding) option -> (Proof.context -> tactic) -> | 
| 120 | local_theory -> (string * Typedef.info) * local_theory | |
| 48975 
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 | structure BNF_Util : BNF_UTIL = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 124 | struct | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 125 | |
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: 
53888diff
changeset | 126 | open Ctr_Sugar_Util | 
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 127 | open BNF_FP_Rec_Sugar_Util | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 128 | |
| 69593 | 129 | val transfer_plugin = Plugin_Name.declare_setup \<^binding>\<open>transfer\<close>; | 
| 64413 | 130 | |
| 131 | ||
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: 
53888diff
changeset | 132 | (* Library proper *) | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 133 | |
| 63797 | 134 | fun unfla0 xs = fold_map (fn _ => fn (c :: cs) => (c, cs)) xs; | 
| 135 | fun unflat0 xss = fold_map unfla0 xss; | |
| 136 | fun unflatt0 xsss = fold_map unflat0 xsss; | |
| 137 | fun unflattt0 xssss = fold_map unflatt0 xssss; | |
| 138 | ||
| 139 | fun unflatt xsss = fst o unflatt0 xsss; | |
| 140 | fun unflattt xssss = fst o unflattt0 xssss; | |
| 141 | ||
| 54601 
91a1e4aa7c80
command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
 traytel parents: 
54544diff
changeset | 142 | val parse_type_arg_constrained = | 
| 69593 | 143 | Parse.type_ident -- Scan.option (\<^keyword>\<open>::\<close> |-- Parse.!!! Parse.sort); | 
| 54601 
91a1e4aa7c80
command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
 traytel parents: 
54544diff
changeset | 144 | |
| 
91a1e4aa7c80
command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
 traytel parents: 
54544diff
changeset | 145 | val parse_type_arg_named_constrained = | 
| 57092 | 146 | (Parse.reserved "dead" >> K NONE || parse_opt_binding_colon >> SOME) -- | 
| 54601 
91a1e4aa7c80
command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
 traytel parents: 
54544diff
changeset | 147 | parse_type_arg_constrained; | 
| 
91a1e4aa7c80
command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
 traytel parents: 
54544diff
changeset | 148 | |
| 
91a1e4aa7c80
command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
 traytel parents: 
54544diff
changeset | 149 | val parse_type_args_named_constrained = | 
| 
91a1e4aa7c80
command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
 traytel parents: 
54544diff
changeset | 150 | parse_type_arg_constrained >> (single o pair (SOME Binding.empty)) || | 
| 69593 | 151 | \<^keyword>\<open>(\<close> |-- Parse.!!! (Parse.list1 parse_type_arg_named_constrained --| \<^keyword>\<open>)\<close>) || | 
| 54601 
91a1e4aa7c80
command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
 traytel parents: 
54544diff
changeset | 152 | Scan.succeed []; | 
| 
91a1e4aa7c80
command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
 traytel parents: 
54544diff
changeset | 153 | |
| 69593 | 154 | val parse_map_rel_pred_binding = Parse.name --| \<^keyword>\<open>:\<close> -- Parse.binding; | 
| 54601 
91a1e4aa7c80
command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
 traytel parents: 
54544diff
changeset | 155 | |
| 62324 | 156 | val no_map_rel = (Binding.empty, Binding.empty, Binding.empty); | 
| 54601 
91a1e4aa7c80
command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
 traytel parents: 
54544diff
changeset | 157 | |
| 62324 | 158 | fun extract_map_rel_pred ("map", m) = (fn (_, r, p) => (m, r, p))
 | 
| 159 |   | extract_map_rel_pred ("rel", r) = (fn (m, _, p) => (m, r, p))
 | |
| 160 |   | extract_map_rel_pred ("pred", p) = (fn (m, r, _) => (m, r, p))
 | |
| 161 |   | extract_map_rel_pred (s, _) = error ("Unknown label " ^ quote s ^ " (expected \"map\" or \"rel\")");
 | |
| 54601 
91a1e4aa7c80
command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
 traytel parents: 
54544diff
changeset | 162 | |
| 62324 | 163 | val parse_map_rel_pred_bindings = | 
| 69593 | 164 | \<^keyword>\<open>for\<close> |-- Scan.repeat parse_map_rel_pred_binding | 
| 62324 | 165 | >> (fn ps => fold extract_map_rel_pred ps no_map_rel) | 
| 57206 
d9be905d6283
changed syntax of map: and rel: arguments to BNF-based datatypes
 blanchet parents: 
57205diff
changeset | 166 | || Scan.succeed no_map_rel; | 
| 54601 
91a1e4aa7c80
command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
 traytel parents: 
54544diff
changeset | 167 | |
| 53263 | 168 | fun typedef (b, Ts, mx) set opt_morphs tac lthy = | 
| 49228 | 169 | let | 
| 53263 | 170 | (*Work around loss of qualification in "typedef" axioms by replicating it in the name*) | 
| 59858 | 171 | val b' = fold_rev Binding.prefix_name (map (suffix "_" o fst) (Binding.path_of b)) b; | 
| 61110 | 172 | |
| 173 | val default_bindings = Typedef.default_bindings (Binding.concealed b'); | |
| 174 | val bindings = | |
| 175 | (case opt_morphs of | |
| 176 | NONE => default_bindings | |
| 177 | | SOME (Rep_name, Abs_name) => | |
| 178 |          {Rep_name = Binding.concealed Rep_name,
 | |
| 179 | Abs_name = Binding.concealed Abs_name, | |
| 180 | type_definition_name = #type_definition_name default_bindings}); | |
| 181 | ||
| 49228 | 182 | val ((name, info), (lthy, lthy_old)) = | 
| 183 | lthy | |
| 72450 | 184 | |> (snd o Local_Theory.begin_nested) | 
| 61260 | 185 |       |> Typedef.add_typedef {overloaded = false} (b', Ts, mx) set (SOME bindings) tac
 | 
| 72450 | 186 | ||> `Local_Theory.end_nested; | 
| 49228 | 187 | val phi = Proof_Context.export_morphism lthy_old lthy; | 
| 188 | in | |
| 189 | ((name, Typedef.transform_info phi info), lthy) | |
| 190 | end; | |
| 191 | ||
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 192 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 193 | (* Term construction *) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 194 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 195 | (** Fresh variables **) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 196 | |
| 49425 | 197 | fun nonzero_string_of_int 0 = "" | 
| 198 | | nonzero_string_of_int n = string_of_int n; | |
| 199 | ||
| 49298 
36e551d3af3b
support for sort constraints in new (co)data commands
 blanchet parents: 
49282diff
changeset | 200 | val mk_TFreess = fold_map mk_TFrees; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 201 | |
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58571diff
changeset | 202 | fun mk_Freesss x Tsss = @{fold_map 2} mk_Freess (mk_names (length Tsss) x) Tsss;
 | 
| 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58571diff
changeset | 203 | fun mk_Freessss x Tssss = @{fold_map 2} mk_Freesss (mk_names (length Tssss) x) Tssss;
 | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 204 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 205 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 206 | (** Types **) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 207 | |
| 53036 | 208 | (*maps [T1,...,Tn]--->T to ([T1,T2,...,Tn], T)*) | 
| 49395 
323414474c1f
use strip_typeN in bnf_def (instead of repairing strip_type)
 traytel parents: 
49366diff
changeset | 209 | fun strip_typeN 0 T = ([], T) | 
| 69593 | 210 | | strip_typeN n (Type (\<^type_name>\<open>fun\<close>, [T, T'])) = strip_typeN (n - 1) T' |>> cons T | 
| 49463 | 211 |   | strip_typeN _ T = raise TYPE ("strip_typeN", [T], []);
 | 
| 49395 
323414474c1f
use strip_typeN in bnf_def (instead of repairing strip_type)
 traytel parents: 
49366diff
changeset | 212 | |
| 53036 | 213 | (*maps [T1,...,Tn]--->T-->U to ([T1,T2,...,Tn], T-->U), where U is not a function type*) | 
| 214 | fun strip_fun_type T = strip_typeN (num_binder_types T - 1) T; | |
| 215 | ||
| 216 | val binder_fun_types = fst o strip_fun_type; | |
| 217 | val body_fun_type = snd o strip_fun_type; | |
| 218 | ||
| 49463 | 219 | fun mk_pred2T T U = mk_predT [T, U]; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 220 | val mk_relT = HOLogic.mk_setT o HOLogic.mk_prodT; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 221 | val dest_relT = HOLogic.dest_prodT o HOLogic.dest_setT; | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 222 | val dest_pred2T = apsnd Term.domain_type o Term.dest_funT; | 
| 69593 | 223 | fun mk_sumT (LT, RT) = Type (\<^type_name>\<open>Sum_Type.sum\<close>, [LT, RT]); | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 224 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 225 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 226 | (** Constants **) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 227 | |
| 69593 | 228 | fun fst_const T = Const (\<^const_name>\<open>fst\<close>, T --> fst (HOLogic.dest_prodT T)); | 
| 229 | fun snd_const T = Const (\<^const_name>\<open>snd\<close>, T --> snd (HOLogic.dest_prodT T)); | |
| 230 | fun Id_const T = Const (\<^const_name>\<open>Id\<close>, mk_relT (T, T)); | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 231 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 232 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 233 | (** Operators **) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 234 | |
| 63824 | 235 | fun enforce_type ctxt get_T T t = | 
| 236 | Term.subst_TVars (tvar_subst (Proof_Context.theory_of ctxt) [get_T (fastype_of t)] [T]) t; | |
| 237 | ||
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 238 | fun mk_converse R = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 239 | let | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 240 | val RT = dest_relT (fastype_of R); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 241 | val RST = mk_relT (snd RT, fst RT); | 
| 69593 | 242 | in Const (\<^const_name>\<open>converse\<close>, fastype_of R --> RST) $ R end; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 243 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 244 | fun mk_rel_comp (R, S) = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 245 | let | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 246 | val RT = fastype_of R; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 247 | val ST = fastype_of S; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 248 | val RST = mk_relT (fst (dest_relT RT), snd (dest_relT ST)); | 
| 69593 | 249 | in Const (\<^const_name>\<open>relcomp\<close>, RT --> ST --> RST) $ R $ S end; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 250 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 251 | fun mk_Gr A f = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 252 | let val ((AT, BT), FT) = `dest_funT (fastype_of f); | 
| 69593 | 253 | in Const (\<^const_name>\<open>Gr\<close>, HOLogic.mk_setT AT --> FT --> mk_relT (AT, BT)) $ A $ f end; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 254 | |
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 255 | fun mk_conversep R = | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 256 | let | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 257 | val RT = dest_pred2T (fastype_of R); | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 258 | val RST = mk_pred2T (snd RT) (fst RT); | 
| 69593 | 259 | in Const (\<^const_name>\<open>conversep\<close>, fastype_of R --> RST) $ R end; | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 260 | |
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 261 | fun mk_rel_compp (R, S) = | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 262 | let | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 263 | val RT = fastype_of R; | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 264 | val ST = fastype_of S; | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 265 | val RST = mk_pred2T (fst (dest_pred2T RT)) (snd (dest_pred2T ST)); | 
| 69593 | 266 | in Const (\<^const_name>\<open>relcompp\<close>, RT --> ST --> RST) $ R $ S end; | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 267 | |
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 268 | fun mk_Grp A f = | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 269 | let val ((AT, BT), FT) = `dest_funT (fastype_of f); | 
| 69593 | 270 | in Const (\<^const_name>\<open>Grp\<close>, HOLogic.mk_setT AT --> FT --> mk_pred2T AT BT) $ A $ f end; | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 271 | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 272 | fun mk_image f = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 273 | let val (T, U) = dest_funT (fastype_of f); | 
| 69593 | 274 | in Const (\<^const_name>\<open>image\<close>, (T --> U) --> HOLogic.mk_setT T --> HOLogic.mk_setT U) $ f end; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 275 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 276 | fun mk_Ball X f = | 
| 69593 | 277 | Const (\<^const_name>\<open>Ball\<close>, fastype_of X --> fastype_of f --> HOLogic.boolT) $ X $ f; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 278 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 279 | fun mk_Bex X f = | 
| 69593 | 280 | Const (\<^const_name>\<open>Bex\<close>, fastype_of X --> fastype_of f --> HOLogic.boolT) $ X $ f; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 281 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 282 | fun mk_UNION X f = | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
62324diff
changeset | 283 | let | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
62324diff
changeset | 284 | val (T, U) = dest_funT (fastype_of f); | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
62324diff
changeset | 285 | in | 
| 69593 | 286 | Const (\<^const_name>\<open>Sup\<close>, HOLogic.mk_setT U --> U) | 
| 287 | $ (Const (\<^const_name>\<open>image\<close>, (T --> U) --> fastype_of X --> HOLogic.mk_setT U) $ f $ X) | |
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
62324diff
changeset | 288 | end; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 289 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 290 | fun mk_Union T = | 
| 69593 | 291 | Const (\<^const_name>\<open>Sup\<close>, HOLogic.mk_setT (HOLogic.mk_setT T) --> HOLogic.mk_setT T); | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 292 | |
| 69593 | 293 | val mk_union = HOLogic.mk_binop \<^const_name>\<open>sup\<close>; | 
| 63851 | 294 | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 295 | fun mk_Field r = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 296 | let val T = fst (dest_relT (fastype_of r)); | 
| 69593 | 297 | in Const (\<^const_name>\<open>Field\<close>, mk_relT (T, T) --> HOLogic.mk_setT T) $ r end; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 298 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 299 | fun mk_card_order bd = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 300 | let | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 301 | val T = fastype_of bd; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 302 | val AT = fst (dest_relT T); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 303 | in | 
| 69593 | 304 | Const (\<^const_name>\<open>card_order_on\<close>, HOLogic.mk_setT AT --> T --> HOLogic.boolT) $ | 
| 55571 | 305 | HOLogic.mk_UNIV AT $ bd | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 306 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 307 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 308 | fun mk_Card_order bd = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 309 | let | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 310 | val T = fastype_of bd; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 311 | val AT = fst (dest_relT T); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 312 | in | 
| 69593 | 313 | Const (\<^const_name>\<open>card_order_on\<close>, HOLogic.mk_setT AT --> T --> HOLogic.boolT) $ | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 314 | mk_Field bd $ bd | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 315 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 316 | |
| 69593 | 317 | fun mk_cinfinite bd = Const (\<^const_name>\<open>cinfinite\<close>, fastype_of bd --> HOLogic.boolT) $ bd; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 318 | |
| 75624 | 319 | fun mk_regularCard bd = Const (\<^const_name>\<open>regularCard\<close>, fastype_of bd --> HOLogic.boolT) $ bd; | 
| 320 | ||
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 321 | fun mk_ordLeq t1 t2 = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 322 | HOLogic.mk_mem (HOLogic.mk_prod (t1, t2), | 
| 69593 | 323 | Const (\<^const_name>\<open>ordLeq\<close>, mk_relT (fastype_of t1, fastype_of t2))); | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 324 | |
| 75624 | 325 | fun mk_ordLess t1 t2 = | 
| 326 | HOLogic.mk_mem (HOLogic.mk_prod (t1, t2), | |
| 327 | Const (\<^const_name>\<open>ordLess\<close>, mk_relT (fastype_of t1, fastype_of t2))); | |
| 328 | ||
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 329 | fun mk_card_of A = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 330 | let | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 331 | val AT = fastype_of A; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 332 | val T = HOLogic.dest_setT AT; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 333 | in | 
| 69593 | 334 | Const (\<^const_name>\<open>card_of\<close>, AT --> mk_relT (T, T)) $ A | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 335 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 336 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 337 | fun mk_dir_image r f = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 338 | let val (T, U) = dest_funT (fastype_of f); | 
| 69593 | 339 | in Const (\<^const_name>\<open>dir_image\<close>, mk_relT (T, T) --> (T --> U) --> mk_relT (U, U)) $ r $ f end; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 340 | |
| 55945 | 341 | fun mk_rel_fun R S = | 
| 52719 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 traytel parents: 
52545diff
changeset | 342 | let | 
| 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 traytel parents: 
52545diff
changeset | 343 | val ((RA, RB), RT) = `dest_pred2T (fastype_of R); | 
| 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 traytel parents: 
52545diff
changeset | 344 | val ((SA, SB), ST) = `dest_pred2T (fastype_of S); | 
| 69593 | 345 | in Const (\<^const_name>\<open>rel_fun\<close>, RT --> ST --> mk_pred2T (RA --> SA) (RB --> SB)) $ R $ S end; | 
| 52719 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 traytel parents: 
52545diff
changeset | 346 | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 347 | (*FIXME: "x"?*) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 348 | (*(nth sets i) must be of type "T --> 'ai set"*) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 349 | fun mk_in As sets T = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 350 | let | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 351 | fun in_single set A = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 352 | let val AT = fastype_of A; | 
| 69593 | 353 |       in Const (\<^const_name>\<open>less_eq\<close>, AT --> AT --> HOLogic.boolT) $ (set $ Free ("x", T)) $ A end;
 | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 354 | in | 
| 58208 | 355 | if null sets then HOLogic.mk_UNIV T | 
| 356 |     else HOLogic.mk_Collect ("x", T, foldr1 (HOLogic.mk_conj) (map2 in_single sets As))
 | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 357 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 358 | |
| 56635 | 359 | fun mk_inj t = | 
| 69593 | 360 | let val T as Type (\<^type_name>\<open>fun\<close>, [domT, _]) = fastype_of t in | 
| 361 | Const (\<^const_name>\<open>inj_on\<close>, T --> HOLogic.mk_setT domT --> HOLogic.boolT) $ t | |
| 56635 | 362 | $ HOLogic.mk_UNIV domT | 
| 363 | end; | |
| 364 | ||
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 365 | fun mk_leq t1 t2 = | 
| 69593 | 366 | Const (\<^const_name>\<open>less_eq\<close>, fastype_of t1 --> fastype_of t2 --> HOLogic.boolT) $ t1 $ t2; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 367 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 368 | fun mk_card_binop binop typop t1 t2 = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 369 | let | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 370 | val (T1, relT1) = `(fst o dest_relT) (fastype_of t1); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 371 | val (T2, relT2) = `(fst o dest_relT) (fastype_of t2); | 
| 55573 | 372 | in Const (binop, relT1 --> relT2 --> mk_relT (typop (T1, T2), typop (T1, T2))) $ t1 $ t2 end; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 373 | |
| 69593 | 374 | val mk_csum = mk_card_binop \<^const_name>\<open>csum\<close> mk_sumT; | 
| 375 | val mk_cprod = mk_card_binop \<^const_name>\<open>cprod\<close> HOLogic.mk_prodT; | |
| 376 | val mk_cexp = mk_card_binop \<^const_name>\<open>cexp\<close> (op --> o swap); | |
| 377 | val ctwo = \<^term>\<open>ctwo\<close>; | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 378 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 379 | fun mk_collect xs defT = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 380 | let val T = (case xs of [] => defT | (x::_) => fastype_of x); | 
| 69593 | 381 | in Const (\<^const_name>\<open>collect\<close>, HOLogic.mk_setT T --> T) $ (HOLogic.mk_set T xs) end; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 382 | |
| 55803 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 traytel parents: 
55772diff
changeset | 383 | fun mk_vimage2p f g = | 
| 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 traytel parents: 
55772diff
changeset | 384 | let | 
| 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 traytel parents: 
55772diff
changeset | 385 | val (T1, T2) = dest_funT (fastype_of f); | 
| 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 traytel parents: 
55772diff
changeset | 386 | val (U1, U2) = dest_funT (fastype_of g); | 
| 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 traytel parents: 
55772diff
changeset | 387 | in | 
| 69593 | 388 | Const (\<^const_name>\<open>vimage2p\<close>, | 
| 55803 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 traytel parents: 
55772diff
changeset | 389 | (T1 --> T2) --> (U1 --> U2) --> mk_pred2T T2 U2 --> mk_pred2T T1 U1) $ f $ g | 
| 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 traytel parents: 
55772diff
changeset | 390 | end; | 
| 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 traytel parents: 
55772diff
changeset | 391 | |
| 62324 | 392 | fun mk_eq_onp P = | 
| 393 | let | |
| 394 | val T = domain_type (fastype_of P); | |
| 395 | in | |
| 69593 | 396 | Const (\<^const_name>\<open>eq_onp\<close>, (T --> HOLogic.boolT) --> T --> T --> HOLogic.boolT) $ P | 
| 62324 | 397 | end; | 
| 398 | ||
| 61240 | 399 | fun mk_pred name R = | 
| 400 | Const (name, uncurry mk_pred2T (fastype_of R |> dest_pred2T) --> HOLogic.boolT) $ R; | |
| 75503 
e5d88927e017
introduced predicate reflp_on and redefined reflp to be an abbreviation
 desharna parents: 
72450diff
changeset | 401 | val mk_reflp = mk_pred \<^const_abbrev>\<open>reflp\<close>; | 
| 76644 
99d6e9217586
added predicates sym_on and symp_on and redefined sym and symp to be abbreviations
 desharna parents: 
75624diff
changeset | 402 | val mk_symp = mk_pred \<^const_abbrev>\<open>symp\<close>; | 
| 76743 
d33fc5228aae
added predicates trans_on and transp_on and redefined trans and transp to be abbreviations
 desharna parents: 
76644diff
changeset | 403 | val mk_transp = mk_pred \<^const_abbrev>\<open>transp\<close>; | 
| 61240 | 404 | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 405 | fun mk_trans thm1 thm2 = trans OF [thm1, thm2]; | 
| 52913 
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
 traytel parents: 
52719diff
changeset | 406 | fun mk_sym thm = thm RS sym; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 407 | |
| 56765 | 408 | val prod_injectD = @{thm iffD1[OF prod.inject]};
 | 
| 409 | val prod_injectI = @{thm iffD2[OF prod.inject]};
 | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 410 | val ctrans = @{thm ordLeq_transitive};
 | 
| 49585 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 blanchet parents: 
49536diff
changeset | 411 | val id_apply = @{thm id_apply};
 | 
| 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 blanchet parents: 
49536diff
changeset | 412 | val meta_mp = @{thm meta_mp};
 | 
| 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 blanchet parents: 
49536diff
changeset | 413 | val meta_spec = @{thm meta_spec};
 | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 414 | val o_apply = @{thm o_apply};
 | 
| 58446 | 415 | val rel_funD = @{thm rel_funD};
 | 
| 416 | val rel_funI = @{thm rel_funI};
 | |
| 49488 
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
 traytel parents: 
49484diff
changeset | 417 | val set_mp = @{thm set_mp};
 | 
| 
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
 traytel parents: 
49484diff
changeset | 418 | val set_rev_mp = @{thm set_rev_mp};
 | 
| 49490 | 419 | val subset_UNIV = @{thm subset_UNIV};
 | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 420 | |
| 62124 | 421 | fun mk_pointful ctxt thm = unfold_thms ctxt [o_apply] (thm RS fun_cong); | 
| 422 | ||
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 423 | fun mk_nthN 1 t 1 = t | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 424 | | mk_nthN _ t 1 = HOLogic.mk_fst t | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 425 | | mk_nthN 2 t 2 = HOLogic.mk_snd t | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 426 | | mk_nthN n t m = mk_nthN (n - 1) (HOLogic.mk_snd t) (m - 1); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 427 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 428 | fun mk_nth_conv n m = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 429 | let | 
| 63798 | 430 |     fun thm b = if b then @{thm fstI} else @{thm sndI};
 | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 431 | fun mk_nth_conv _ 1 1 = refl | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 432 |       | mk_nth_conv _ _ 1 = @{thm fst_conv}
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 433 |       | mk_nth_conv _ 2 2 = @{thm snd_conv}
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 434 |       | mk_nth_conv b _ 2 = @{thm snd_conv} RS thm b
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 435 | | mk_nth_conv b n m = mk_nth_conv false (n - 1) (m - 1) RS thm b; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 436 | in mk_nth_conv (not (m = n)) n m end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 437 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 438 | fun mk_nthI 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 | 439 |   | mk_nthI n m = fold (curry op RS) (replicate (m - 1) @{thm sndI})
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 440 |     (if m = n then @{thm TrueE[OF TrueI]} else @{thm fstI});
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 441 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 442 | fun mk_conjIN 1 = @{thm TrueE[OF TrueI]}
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 443 | | mk_conjIN n = mk_conjIN (n - 1) RSN (2, conjI); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 444 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 445 | fun mk_ordLeq_csum 1 1 thm = thm | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 446 |   | mk_ordLeq_csum _ 1 thm = @{thm ordLeq_transitive} OF [thm, @{thm ordLeq_csum1}]
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 447 |   | mk_ordLeq_csum 2 2 thm = @{thm ordLeq_transitive} OF [thm, @{thm ordLeq_csum2}]
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 448 |   | mk_ordLeq_csum n m thm = @{thm ordLeq_transitive} OF
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 449 |     [mk_ordLeq_csum (n - 1) (m - 1) thm, @{thm ordLeq_csum2[OF Card_order_csum]}];
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 450 | |
| 75624 | 451 | fun mk_ordLess_csum 1 1 thm = thm | 
| 452 |   | mk_ordLess_csum _ 1 thm = @{thm ordLess_ordLeq_trans} OF [thm, @{thm ordLeq_csum1}]
 | |
| 453 |   | mk_ordLess_csum 2 2 thm = @{thm ordLess_ordLeq_trans} OF [thm, @{thm ordLeq_csum2}]
 | |
| 454 |   | mk_ordLess_csum n m thm = @{thm ordLess_ordLeq_trans} OF
 | |
| 455 |     [mk_ordLess_csum (n - 1) (m - 1) thm, @{thm ordLeq_csum2[OF Card_order_csum]}];
 | |
| 456 | ||
| 58676 | 457 | fun mk_rel_funDN n = funpow n (fn thm => thm RS rel_funD); | 
| 58446 | 458 | |
| 459 | val mk_rel_funDN_rotated = rotate_prems ~1 oo mk_rel_funDN; | |
| 460 | ||
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 461 | local | 
| 63399 | 462 |   fun mk_Un_upper' 0 = @{thm subset_refl}
 | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 463 |     | mk_Un_upper' 1 = @{thm Un_upper1}
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 464 | | mk_Un_upper' k = Library.foldr (op RS o swap) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 465 |       (replicate (k - 1) @{thm subset_trans[OF Un_upper1]}, @{thm Un_upper1});
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 466 | in | 
| 63399 | 467 |   fun mk_Un_upper 1 1 = @{thm subset_refl}
 | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 468 |     | mk_Un_upper n 1 = mk_Un_upper' (n - 2) RS @{thm subset_trans[OF Un_upper1]}
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 469 |     | mk_Un_upper n m = mk_Un_upper' (n - m) RS @{thm subset_trans[OF Un_upper2]};
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 470 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 471 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 472 | local | 
| 49366 | 473 |   fun mk_UnIN' 0 = @{thm UnI2}
 | 
| 474 |     | mk_UnIN' m = mk_UnIN' (m - 1) RS @{thm UnI1};
 | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 475 | in | 
| 49366 | 476 |   fun mk_UnIN 1 1 = @{thm TrueE[OF TrueI]}
 | 
| 477 |     | mk_UnIN n 1 = Library.foldr1 (op RS o swap) (replicate (n - 1) @{thm UnI1})
 | |
| 478 | | mk_UnIN n m = mk_UnIN' (n - m) | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 479 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 480 | |
| 57205 | 481 | fun is_refl_bool t = | 
| 482 | op aconv (HOLogic.dest_eq t) | |
| 483 | handle TERM _ => false; | |
| 484 | ||
| 49605 
ea566f5e1724
avoid another brand of trivial "disc_rel" theorems (which made the simplifier loop for all single-constructor types)
 blanchet parents: 
49585diff
changeset | 485 | fun is_refl_prop t = | 
| 
ea566f5e1724
avoid another brand of trivial "disc_rel" theorems (which made the simplifier loop for all single-constructor types)
 blanchet parents: 
49585diff
changeset | 486 | op aconv (HOLogic.dest_eq (HOLogic.dest_Trueprop t)) | 
| 
ea566f5e1724
avoid another brand of trivial "disc_rel" theorems (which made the simplifier loop for all single-constructor types)
 blanchet parents: 
49585diff
changeset | 487 | handle TERM _ => false; | 
| 
ea566f5e1724
avoid another brand of trivial "disc_rel" theorems (which made the simplifier loop for all single-constructor types)
 blanchet parents: 
49585diff
changeset | 488 | |
| 
ea566f5e1724
avoid another brand of trivial "disc_rel" theorems (which made the simplifier loop for all single-constructor types)
 blanchet parents: 
49585diff
changeset | 489 | val is_refl = is_refl_prop o Thm.prop_of; | 
| 
ea566f5e1724
avoid another brand of trivial "disc_rel" theorems (which made the simplifier loop for all single-constructor types)
 blanchet parents: 
49585diff
changeset | 490 | val is_concl_refl = is_refl_prop o Logic.strip_imp_concl o Thm.prop_of; | 
| 
ea566f5e1724
avoid another brand of trivial "disc_rel" theorems (which made the simplifier loop for all single-constructor types)
 blanchet parents: 
49585diff
changeset | 491 | |
| 49484 | 492 | val no_refl = filter_out is_refl; | 
| 493 | val no_reflexive = filter_out Thm.is_reflexive; | |
| 494 | ||
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 495 | end; |