| author | wenzelm | 
| Sun, 26 Nov 2017 21:08:32 +0100 | |
| changeset 67091 | 1393c2340eec | 
| parent 64413 | c0d5e78eb647 | 
| child 69593 | 3dda49e08b9d | 
| 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 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 56 | 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 | 57 | val mk_converse: term -> term | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 58 | 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 | 59 | 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 | 60 | 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 | 61 | val mk_dir_image: term -> term -> term | 
| 62324 | 62 | val mk_eq_onp: term -> term | 
| 55945 | 63 | 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 | 64 | val mk_image: term -> term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 65 | val mk_in: term list -> term list -> typ -> term | 
| 56635 | 66 | val mk_inj: term -> term | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 67 | 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 | 68 | val mk_ordLeq: term -> term -> term | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 69 | 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 | 70 | 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 | 71 | val mk_vimage2p: term -> term -> term | 
| 61240 | 72 | val mk_reflp: term -> term | 
| 73 | val mk_symp: term -> term | |
| 74 | val mk_transp: term -> term | |
| 63851 | 75 | 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 | 76 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 77 | (*parameterized terms*) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 78 | 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 | 79 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 80 | (*parameterized thms*) | 
| 56765 | 81 | val prod_injectD: thm | 
| 82 | val prod_injectI: thm | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 83 | val ctrans: thm | 
| 49585 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 blanchet parents: 
49536diff
changeset | 84 | val id_apply: thm | 
| 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 blanchet parents: 
49536diff
changeset | 85 | val meta_mp: thm | 
| 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 blanchet parents: 
49536diff
changeset | 86 | val meta_spec: thm | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 87 | val o_apply: thm | 
| 58446 | 88 | val rel_funD: thm | 
| 89 | val rel_funI: thm | |
| 49488 
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
 traytel parents: 
49484diff
changeset | 90 | val set_mp: thm | 
| 
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
 traytel parents: 
49484diff
changeset | 91 | val set_rev_mp: thm | 
| 49490 | 92 | val subset_UNIV: thm | 
| 58446 | 93 | |
| 94 | val mk_conjIN: int -> thm | |
| 95 | val mk_nthI: int -> int -> thm | |
| 96 | val mk_nth_conv: int -> int -> thm | |
| 97 | val mk_ordLeq_csum: int -> int -> thm -> thm | |
| 62124 | 98 | val mk_pointful: Proof.context -> thm -> thm | 
| 58446 | 99 | val mk_rel_funDN: int -> thm -> thm | 
| 100 | 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 | 101 | val mk_sym: thm -> thm | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 102 | val mk_trans: thm -> thm -> thm | 
| 58446 | 103 | val mk_UnIN: int -> int -> thm | 
| 104 | 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 | 105 | |
| 57205 | 106 | val is_refl_bool: term -> bool | 
| 49484 | 107 | 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 | 108 | val is_concl_refl: thm -> bool | 
| 49484 | 109 | val no_refl: thm list -> thm list | 
| 110 | val no_reflexive: thm list -> thm list | |
| 111 | ||
| 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 | 112 | val parse_type_args_named_constrained: (binding option * (string * string option)) list parser | 
| 62324 | 113 | 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 | 114 | |
| 49835 | 115 | val typedef: binding * (string * sort) list * mixfix -> term -> | 
| 58959 | 116 | (binding * binding) option -> (Proof.context -> tactic) -> | 
| 117 | 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 | 118 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 119 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 120 | structure BNF_Util : BNF_UTIL = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 121 | struct | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 122 | |
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: 
53888diff
changeset | 123 | open Ctr_Sugar_Util | 
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55573diff
changeset | 124 | 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 | 125 | |
| 64413 | 126 | val transfer_plugin = Plugin_Name.declare_setup @{binding transfer};
 | 
| 127 | ||
| 128 | ||
| 54008 
b15cfc2864de
refactoring -- splitting between constructor sugar dependencies and true BNF dependencies
 blanchet parents: 
53888diff
changeset | 129 | (* Library proper *) | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 130 | |
| 63797 | 131 | fun unfla0 xs = fold_map (fn _ => fn (c :: cs) => (c, cs)) xs; | 
| 132 | fun unflat0 xss = fold_map unfla0 xss; | |
| 133 | fun unflatt0 xsss = fold_map unflat0 xsss; | |
| 134 | fun unflattt0 xssss = fold_map unflatt0 xssss; | |
| 135 | ||
| 136 | fun unflatt xsss = fst o unflatt0 xsss; | |
| 137 | fun unflattt xssss = fst o unflattt0 xssss; | |
| 138 | ||
| 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 | 139 | val 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 | 140 |   Parse.type_ident -- Scan.option (@{keyword "::"} |-- Parse.!!! Parse.sort);
 | 
| 
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 | 141 | |
| 
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_named_constrained = | 
| 57092 | 143 | (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 | 144 | 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 | 145 | |
| 
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 | 146 | 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 | 147 | parse_type_arg_constrained >> (single o pair (SOME Binding.empty)) || | 
| 
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 |   @{keyword "("} |-- Parse.!!! (Parse.list1 parse_type_arg_named_constrained --| @{keyword ")"}) ||
 | 
| 
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 | 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 | 150 | |
| 62324 | 151 | val parse_map_rel_pred_binding = Parse.name --| @{keyword ":"} -- 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 | 152 | |
| 62324 | 153 | 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 | 154 | |
| 62324 | 155 | fun extract_map_rel_pred ("map", m) = (fn (_, r, p) => (m, r, p))
 | 
| 156 |   | extract_map_rel_pred ("rel", r) = (fn (m, _, p) => (m, r, p))
 | |
| 157 |   | extract_map_rel_pred ("pred", p) = (fn (m, r, _) => (m, r, p))
 | |
| 158 |   | 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 | 159 | |
| 62324 | 160 | val parse_map_rel_pred_bindings = | 
| 161 |   @{keyword "for"} |-- Scan.repeat parse_map_rel_pred_binding
 | |
| 162 | >> (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 | 163 | || 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 | 164 | |
| 53263 | 165 | fun typedef (b, Ts, mx) set opt_morphs tac lthy = | 
| 49228 | 166 | let | 
| 53263 | 167 | (*Work around loss of qualification in "typedef" axioms by replicating it in the name*) | 
| 59858 | 168 | val b' = fold_rev Binding.prefix_name (map (suffix "_" o fst) (Binding.path_of b)) b; | 
| 61110 | 169 | |
| 170 | val default_bindings = Typedef.default_bindings (Binding.concealed b'); | |
| 171 | val bindings = | |
| 172 | (case opt_morphs of | |
| 173 | NONE => default_bindings | |
| 174 | | SOME (Rep_name, Abs_name) => | |
| 175 |          {Rep_name = Binding.concealed Rep_name,
 | |
| 176 | Abs_name = Binding.concealed Abs_name, | |
| 177 | type_definition_name = #type_definition_name default_bindings}); | |
| 178 | ||
| 49228 | 179 | val ((name, info), (lthy, lthy_old)) = | 
| 180 | lthy | |
| 61101 
7b915ca69af1
use open/close_target rather than Local_Theory.restore to get polymorphic definitions;
 traytel parents: 
59880diff
changeset | 181 | |> Local_Theory.open_target |> snd | 
| 61260 | 182 |       |> Typedef.add_typedef {overloaded = false} (b', Ts, mx) set (SOME bindings) tac
 | 
| 61101 
7b915ca69af1
use open/close_target rather than Local_Theory.restore to get polymorphic definitions;
 traytel parents: 
59880diff
changeset | 183 | ||> `Local_Theory.close_target; | 
| 49228 | 184 | val phi = Proof_Context.export_morphism lthy_old lthy; | 
| 185 | in | |
| 186 | ((name, Typedef.transform_info phi info), lthy) | |
| 187 | end; | |
| 188 | ||
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 189 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 190 | (* Term construction *) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 191 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 192 | (** Fresh variables **) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 193 | |
| 49425 | 194 | fun nonzero_string_of_int 0 = "" | 
| 195 | | nonzero_string_of_int n = string_of_int n; | |
| 196 | ||
| 49298 
36e551d3af3b
support for sort constraints in new (co)data commands
 blanchet parents: 
49282diff
changeset | 197 | 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 | 198 | |
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58571diff
changeset | 199 | 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 | 200 | 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 | 201 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 202 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 203 | (** Types **) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 204 | |
| 53036 | 205 | (*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 | 206 | fun strip_typeN 0 T = ([], T) | 
| 
323414474c1f
use strip_typeN in bnf_def (instead of repairing strip_type)
 traytel parents: 
49366diff
changeset | 207 |   | strip_typeN n (Type (@{type_name fun}, [T, T'])) = strip_typeN (n - 1) T' |>> cons T
 | 
| 49463 | 208 |   | 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 | 209 | |
| 53036 | 210 | (*maps [T1,...,Tn]--->T-->U to ([T1,T2,...,Tn], T-->U), where U is not a function type*) | 
| 211 | fun strip_fun_type T = strip_typeN (num_binder_types T - 1) T; | |
| 212 | ||
| 213 | val binder_fun_types = fst o strip_fun_type; | |
| 214 | val body_fun_type = snd o strip_fun_type; | |
| 215 | ||
| 49463 | 216 | 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 | 217 | 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 | 218 | 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 | 219 | val dest_pred2T = apsnd Term.domain_type o Term.dest_funT; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 220 | fun mk_sumT (LT, RT) = Type (@{type_name Sum_Type.sum}, [LT, RT]);
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 221 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 222 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 223 | (** Constants **) | 
| 
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 | fun fst_const T = Const (@{const_name fst}, T --> fst (HOLogic.dest_prodT T));
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 226 | fun snd_const T = Const (@{const_name snd}, T --> snd (HOLogic.dest_prodT T));
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 227 | fun Id_const T = Const (@{const_name Id}, mk_relT (T, T));
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 228 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 229 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 230 | (** Operators **) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 231 | |
| 63824 | 232 | fun enforce_type ctxt get_T T t = | 
| 233 | Term.subst_TVars (tvar_subst (Proof_Context.theory_of ctxt) [get_T (fastype_of t)] [T]) t; | |
| 234 | ||
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 235 | fun mk_converse R = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 236 | let | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 237 | 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 | 238 | val RST = mk_relT (snd RT, fst RT); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 239 |   in Const (@{const_name converse}, fastype_of R --> RST) $ R end;
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 240 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 241 | 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 | 242 | let | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 243 | val RT = fastype_of R; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 244 | val ST = fastype_of S; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 245 | val RST = mk_relT (fst (dest_relT RT), snd (dest_relT ST)); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 246 |   in Const (@{const_name relcomp}, RT --> ST --> RST) $ R $ S end;
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 247 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 248 | fun mk_Gr A f = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 249 | let val ((AT, BT), FT) = `dest_funT (fastype_of f); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 250 |   in Const (@{const_name Gr}, HOLogic.mk_setT AT --> FT --> mk_relT (AT, BT)) $ A $ f end;
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 251 | |
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 252 | fun mk_conversep R = | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 253 | let | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 254 | 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 | 255 | val RST = mk_pred2T (snd RT) (fst RT); | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 256 |   in Const (@{const_name conversep}, fastype_of R --> RST) $ R end;
 | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 257 | |
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 258 | fun mk_rel_compp (R, S) = | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 259 | let | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 260 | val RT = fastype_of R; | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 261 | val ST = fastype_of S; | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 262 | val RST = mk_pred2T (fst (dest_pred2T RT)) (snd (dest_pred2T ST)); | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 263 |   in Const (@{const_name relcompp}, RT --> ST --> RST) $ R $ S end;
 | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 264 | |
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 265 | fun mk_Grp A f = | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 266 | let val ((AT, BT), FT) = `dest_funT (fastype_of f); | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 267 |   in Const (@{const_name Grp}, HOLogic.mk_setT AT --> FT --> mk_pred2T AT BT) $ A $ f end;
 | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 268 | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 269 | fun mk_image f = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 270 | let val (T, U) = dest_funT (fastype_of f); | 
| 55573 | 271 |   in Const (@{const_name image}, (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 | 272 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 273 | fun mk_Ball X f = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 274 |   Const (@{const_name Ball}, fastype_of X --> fastype_of f --> HOLogic.boolT) $ X $ f;
 | 
| 
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_Bex X f = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 277 |   Const (@{const_name Bex}, fastype_of X --> fastype_of f --> HOLogic.boolT) $ X $ f;
 | 
| 
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_UNION X f = | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
62324diff
changeset | 280 | let | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
62324diff
changeset | 281 | val (T, U) = dest_funT (fastype_of f); | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
62324diff
changeset | 282 | in | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
62324diff
changeset | 283 |     Const (@{const_name Sup}, HOLogic.mk_setT U --> U)
 | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
62324diff
changeset | 284 |       $ (Const (@{const_name image}, (T --> U) --> fastype_of X --> HOLogic.mk_setT U) $ f $ X)
 | 
| 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
62324diff
changeset | 285 | end; | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 286 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 287 | fun mk_Union T = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 288 |   Const (@{const_name Sup}, HOLogic.mk_setT (HOLogic.mk_setT T) --> HOLogic.mk_setT T);
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 289 | |
| 63851 | 290 | val mk_union = HOLogic.mk_binop @{const_name sup};
 | 
| 291 | ||
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 292 | fun mk_Field r = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 293 | let val T = fst (dest_relT (fastype_of r)); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 294 |   in Const (@{const_name Field}, mk_relT (T, T) --> HOLogic.mk_setT T) $ r end;
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 295 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 296 | fun mk_card_order bd = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 297 | let | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 298 | val T = fastype_of bd; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 299 | 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 | 300 | in | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 301 |     Const (@{const_name card_order_on}, HOLogic.mk_setT AT --> T --> HOLogic.boolT) $
 | 
| 55571 | 302 | 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 | 303 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 304 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 305 | fun mk_Card_order bd = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 306 | let | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 307 | val T = fastype_of bd; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 308 | 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 | 309 | in | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 310 |     Const (@{const_name card_order_on}, HOLogic.mk_setT AT --> T --> HOLogic.boolT) $
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 311 | mk_Field bd $ bd | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 312 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 313 | |
| 55573 | 314 | fun mk_cinfinite bd = Const (@{const_name cinfinite}, 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 | 315 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 316 | fun mk_ordLeq t1 t2 = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 317 | HOLogic.mk_mem (HOLogic.mk_prod (t1, t2), | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 318 |     Const (@{const_name ordLeq}, mk_relT (fastype_of t1, fastype_of t2)));
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 319 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 320 | fun mk_card_of A = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 321 | let | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 322 | val AT = fastype_of A; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 323 | 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 | 324 | in | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 325 |     Const (@{const_name card_of}, AT --> mk_relT (T, T)) $ A
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 326 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 327 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 328 | 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 | 329 | let val (T, U) = dest_funT (fastype_of f); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 330 |   in Const (@{const_name dir_image}, mk_relT (T, T) --> (T --> U) --> mk_relT (U, U)) $ r $ f end;
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 331 | |
| 55945 | 332 | fun mk_rel_fun R S = | 
| 52719 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 traytel parents: 
52545diff
changeset | 333 | let | 
| 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 traytel parents: 
52545diff
changeset | 334 | 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 | 335 | val ((SA, SB), ST) = `dest_pred2T (fastype_of S); | 
| 55945 | 336 |   in Const (@{const_name rel_fun}, 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 | 337 | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 338 | (*FIXME: "x"?*) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 339 | (*(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 | 340 | 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 | 341 | let | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 342 | fun in_single set A = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 343 | let val AT = fastype_of A; | 
| 58208 | 344 |       in Const (@{const_name less_eq}, 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 | 345 | in | 
| 58208 | 346 | if null sets then HOLogic.mk_UNIV T | 
| 347 |     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 | 348 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 349 | |
| 56635 | 350 | fun mk_inj t = | 
| 56648 
2ffa440b3074
manual merge + added 'rel_distincts' field to record for symmetry
 blanchet parents: 
56640diff
changeset | 351 |   let val T as Type (@{type_name fun}, [domT, _]) = fastype_of t in
 | 
| 56635 | 352 |     Const (@{const_name inj_on}, T --> HOLogic.mk_setT domT --> HOLogic.boolT) $ t
 | 
| 353 | $ HOLogic.mk_UNIV domT | |
| 354 | end; | |
| 355 | ||
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51861diff
changeset | 356 | fun mk_leq t1 t2 = | 
| 59856 
ed0ca9029021
export more low-level theorems in data structure (partly for 'corec')
 blanchet parents: 
59271diff
changeset | 357 |   Const (@{const_name less_eq}, 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 | 358 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 359 | 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 | 360 | let | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 361 | 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 | 362 | val (T2, relT2) = `(fst o dest_relT) (fastype_of t2); | 
| 55573 | 363 | 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 | 364 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 365 | val mk_csum = mk_card_binop @{const_name csum} mk_sumT;
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 366 | val mk_cprod = mk_card_binop @{const_name cprod} HOLogic.mk_prodT;
 | 
| 52545 
d2ad6eae514f
Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
 traytel parents: 
52280diff
changeset | 367 | val mk_cexp = mk_card_binop @{const_name cexp} (op --> o swap);
 | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 368 | val ctwo = @{term ctwo};
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 369 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 370 | fun mk_collect xs defT = | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 371 | let val T = (case xs of [] => defT | (x::_) => fastype_of x); | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 372 |   in Const (@{const_name collect}, HOLogic.mk_setT T --> T) $ (HOLogic.mk_set T xs) end;
 | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 373 | |
| 55803 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 traytel parents: 
55772diff
changeset | 374 | fun mk_vimage2p f g = | 
| 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 traytel parents: 
55772diff
changeset | 375 | let | 
| 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 traytel parents: 
55772diff
changeset | 376 | 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 | 377 | 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 | 378 | in | 
| 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 traytel parents: 
55772diff
changeset | 379 |     Const (@{const_name vimage2p},
 | 
| 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 traytel parents: 
55772diff
changeset | 380 | (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 | 381 | end; | 
| 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 traytel parents: 
55772diff
changeset | 382 | |
| 62324 | 383 | fun mk_eq_onp P = | 
| 384 | let | |
| 385 | val T = domain_type (fastype_of P); | |
| 386 | in | |
| 387 |     Const (@{const_name eq_onp}, (T --> HOLogic.boolT) --> T --> T --> HOLogic.boolT) $ P
 | |
| 388 | end; | |
| 389 | ||
| 61240 | 390 | fun mk_pred name R = | 
| 391 | Const (name, uncurry mk_pred2T (fastype_of R |> dest_pred2T) --> HOLogic.boolT) $ R; | |
| 392 | val mk_reflp = mk_pred @{const_name reflp};
 | |
| 393 | val mk_symp = mk_pred @{const_name symp};
 | |
| 394 | val mk_transp =  mk_pred @{const_name transp};
 | |
| 395 | ||
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 396 | 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 | 397 | 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 | 398 | |
| 56765 | 399 | val prod_injectD = @{thm iffD1[OF prod.inject]};
 | 
| 400 | 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 | 401 | val ctrans = @{thm ordLeq_transitive};
 | 
| 49585 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 blanchet parents: 
49536diff
changeset | 402 | val id_apply = @{thm id_apply};
 | 
| 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 blanchet parents: 
49536diff
changeset | 403 | val meta_mp = @{thm meta_mp};
 | 
| 
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
 blanchet parents: 
49536diff
changeset | 404 | 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 | 405 | val o_apply = @{thm o_apply};
 | 
| 58446 | 406 | val rel_funD = @{thm rel_funD};
 | 
| 407 | val rel_funI = @{thm rel_funI};
 | |
| 49488 
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
 traytel parents: 
49484diff
changeset | 408 | val set_mp = @{thm set_mp};
 | 
| 
02eb07152998
use iffD* instead of (s)subst instantiated with identity; tuned antiquotations;
 traytel parents: 
49484diff
changeset | 409 | val set_rev_mp = @{thm set_rev_mp};
 | 
| 49490 | 410 | 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 | 411 | |
| 62124 | 412 | fun mk_pointful ctxt thm = unfold_thms ctxt [o_apply] (thm RS fun_cong); | 
| 413 | ||
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 414 | 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 | 415 | | 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 | 416 | | 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 | 417 | | 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 | 418 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 419 | 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 | 420 | let | 
| 63798 | 421 |     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 | 422 | 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 | 423 |       | 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 | 424 |       | 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 | 425 |       | 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 | 426 | | 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 | 427 | 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 | 428 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 429 | 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 | 430 |   | 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 | 431 |     (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 | 432 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 433 | 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 | 434 | | 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 | 435 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 436 | 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 | 437 |   | 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 | 438 |   | 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 | 439 |   | 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 | 440 |     [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 | 441 | |
| 58676 | 442 | fun mk_rel_funDN n = funpow n (fn thm => thm RS rel_funD); | 
| 58446 | 443 | |
| 444 | val mk_rel_funDN_rotated = rotate_prems ~1 oo mk_rel_funDN; | |
| 445 | ||
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 446 | local | 
| 63399 | 447 |   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 | 448 |     | 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 | 449 | | 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 | 450 |       (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 | 451 | in | 
| 63399 | 452 |   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 | 453 |     | 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 | 454 |     | 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 | 455 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 456 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 457 | local | 
| 49366 | 458 |   fun mk_UnIN' 0 = @{thm UnI2}
 | 
| 459 |     | 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 | 460 | in | 
| 49366 | 461 |   fun mk_UnIN 1 1 = @{thm TrueE[OF TrueI]}
 | 
| 462 |     | mk_UnIN n 1 = Library.foldr1 (op RS o swap) (replicate (n - 1) @{thm UnI1})
 | |
| 463 | | 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 | 464 | end; | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 465 | |
| 57205 | 466 | fun is_refl_bool t = | 
| 467 | op aconv (HOLogic.dest_eq t) | |
| 468 | handle TERM _ => false; | |
| 469 | ||
| 49605 
ea566f5e1724
avoid another brand of trivial "disc_rel" theorems (which made the simplifier loop for all single-constructor types)
 blanchet parents: 
49585diff
changeset | 470 | 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 | 471 | 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 | 472 | 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 | 473 | |
| 
ea566f5e1724
avoid another brand of trivial "disc_rel" theorems (which made the simplifier loop for all single-constructor types)
 blanchet parents: 
49585diff
changeset | 474 | 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 | 475 | 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 | 476 | |
| 49484 | 477 | val no_refl = filter_out is_refl; | 
| 478 | val no_reflexive = filter_out Thm.is_reflexive; | |
| 479 | ||
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 480 | end; |