| author | wenzelm | 
| Sun, 23 Feb 2014 15:35:19 +0100 | |
| changeset 55688 | 767edb2c1e4e | 
| parent 55480 | 59cc4a8bc28a | 
| child 55854 | ee270328a781 | 
| permissions | -rw-r--r-- | 
| 55061 | 1  | 
(* Title: HOL/Tools/BNF/bnf_def.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  | 
Author: Jasmin Blanchette, TU Muenchen  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
4  | 
Copyright 2012  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
5  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
6  | 
Definition of bounded natural functors.  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
7  | 
*)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
8  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
9  | 
signature BNF_DEF =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
10  | 
sig  | 
| 
51837
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
11  | 
type bnf  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
12  | 
  type nonemptiness_witness = {I: int list, wit: term, prop: thm list}
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
13  | 
|
| 
51837
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
14  | 
val morph_bnf: morphism -> bnf -> bnf  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
15  | 
val bnf_of: Proof.context -> string -> bnf option  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
16  | 
val register_bnf: string -> (bnf * local_theory) -> (bnf * local_theory)  | 
| 
49434
 
433dc7e028c8
separated registration of BNFs from bnf_def (BNFs are now stored only for bnf_def and (co)data commands)
 
traytel 
parents: 
49430 
diff
changeset
 | 
17  | 
|
| 
51837
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
18  | 
val name_of_bnf: bnf -> binding  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
19  | 
val T_of_bnf: bnf -> typ  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
20  | 
val live_of_bnf: bnf -> int  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
21  | 
val lives_of_bnf: bnf -> typ list  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
22  | 
val dead_of_bnf: bnf -> int  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
23  | 
val deads_of_bnf: bnf -> typ list  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
24  | 
val nwits_of_bnf: bnf -> int  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
25  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
26  | 
val mapN: string  | 
| 49507 | 27  | 
val relN: string  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
28  | 
val setN: string  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
29  | 
val mk_setN: int -> string  | 
| 
55025
 
1ac0a0194428
support declaration of nonemptiness witnesses in bnf_decl
 
traytel 
parents: 
54921 
diff
changeset
 | 
30  | 
val mk_witN: int -> string  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
31  | 
|
| 
51837
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
32  | 
val map_of_bnf: bnf -> term  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
33  | 
val sets_of_bnf: bnf -> term list  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
34  | 
val rel_of_bnf: bnf -> term  | 
| 
49214
 
2a3cb4c71b87
construct the right iterator theorem in the recursive case
 
blanchet 
parents: 
49123 
diff
changeset
 | 
35  | 
|
| 
51837
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
36  | 
val mk_T_of_bnf: typ list -> typ list -> bnf -> typ  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
37  | 
val mk_bd_of_bnf: typ list -> typ list -> bnf -> term  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
38  | 
val mk_map_of_bnf: typ list -> typ list -> typ list -> bnf -> term  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
39  | 
val mk_rel_of_bnf: typ list -> typ list -> typ list -> bnf -> term  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
40  | 
val mk_sets_of_bnf: typ list list -> typ list list -> bnf -> term list  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
41  | 
val mk_wits_of_bnf: typ list list -> typ list list -> bnf -> (int list * term) list  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
42  | 
|
| 
51837
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
43  | 
val bd_Card_order_of_bnf: bnf -> thm  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
44  | 
val bd_Cinfinite_of_bnf: bnf -> thm  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
45  | 
val bd_Cnotzero_of_bnf: bnf -> thm  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
46  | 
val bd_card_order_of_bnf: bnf -> thm  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
47  | 
val bd_cinfinite_of_bnf: bnf -> thm  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
48  | 
val collect_set_map_of_bnf: bnf -> thm  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
49  | 
val in_bd_of_bnf: bnf -> thm  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
50  | 
val in_cong_of_bnf: bnf -> thm  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
51  | 
val in_mono_of_bnf: bnf -> thm  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
52  | 
val in_rel_of_bnf: bnf -> thm  | 
| 53287 | 53  | 
val map_comp0_of_bnf: bnf -> thm  | 
| 53288 | 54  | 
val map_comp_of_bnf: bnf -> thm  | 
| 
51837
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
55  | 
val map_cong0_of_bnf: bnf -> thm  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
56  | 
val map_cong_of_bnf: bnf -> thm  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
57  | 
val map_def_of_bnf: bnf -> thm  | 
| 53270 | 58  | 
val map_id0_of_bnf: bnf -> thm  | 
| 53285 | 59  | 
val map_id_of_bnf: bnf -> thm  | 
| 
52719
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
60  | 
val map_transfer_of_bnf: bnf -> thm  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
61  | 
val le_rel_OO_of_bnf: bnf -> thm  | 
| 
51837
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
62  | 
val rel_def_of_bnf: bnf -> thm  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
63  | 
val rel_Grp_of_bnf: bnf -> thm  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
64  | 
val rel_OO_of_bnf: bnf -> thm  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
65  | 
val rel_OO_Grp_of_bnf: bnf -> thm  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
66  | 
val rel_cong_of_bnf: bnf -> thm  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
67  | 
val rel_conversep_of_bnf: bnf -> thm  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
68  | 
val rel_mono_of_bnf: bnf -> thm  | 
| 51916 | 69  | 
val rel_mono_strong_of_bnf: bnf -> thm  | 
| 
51837
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
70  | 
val rel_eq_of_bnf: bnf -> thm  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
71  | 
val rel_flip_of_bnf: bnf -> thm  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
72  | 
val set_bd_of_bnf: bnf -> thm list  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
73  | 
val set_defs_of_bnf: bnf -> thm list  | 
| 53289 | 74  | 
val set_map0_of_bnf: bnf -> thm list  | 
| 53290 | 75  | 
val set_map_of_bnf: bnf -> thm list  | 
| 
51837
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
76  | 
val wit_thms_of_bnf: bnf -> thm list  | 
| 
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
77  | 
val wit_thmss_of_bnf: bnf -> thm list list  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
78  | 
|
| 54236 | 79  | 
val mk_map: int -> typ list -> typ list -> term -> term  | 
80  | 
val mk_rel: int -> typ list -> typ list -> term -> term  | 
|
| 
54237
 
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
 
blanchet 
parents: 
54236 
diff
changeset
 | 
81  | 
val build_map: Proof.context -> (typ * typ -> term) -> typ * typ -> term  | 
| 
 
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
 
blanchet 
parents: 
54236 
diff
changeset
 | 
82  | 
val build_rel: Proof.context -> (typ * typ -> term) -> typ * typ -> term  | 
| 
54246
 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 
blanchet 
parents: 
54237 
diff
changeset
 | 
83  | 
val flatten_type_args_of_bnf: bnf -> 'a -> 'a list -> 'a list  | 
| 
 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 
blanchet 
parents: 
54237 
diff
changeset
 | 
84  | 
val map_flattened_map_args: Proof.context -> string -> (term list -> 'a list) -> term list ->  | 
| 
 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 
blanchet 
parents: 
54237 
diff
changeset
 | 
85  | 
'a list  | 
| 54236 | 86  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
87  | 
val mk_witness: int list * term -> thm list -> nonemptiness_witness  | 
| 49103 | 88  | 
  val minimize_wits: (''a list * 'b) list -> (''a list * 'b) list
 | 
| 
51837
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
89  | 
val wits_of_bnf: bnf -> nonemptiness_witness list  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
90  | 
|
| 
52635
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
91  | 
val zip_axioms: 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list  | 
| 49456 | 92  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
93  | 
datatype const_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline  | 
| 49538 | 94  | 
datatype fact_policy = Dont_Note | Note_Some | Note_All  | 
95  | 
||
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
96  | 
val bnf_note_all: bool Config.T  | 
| 
53143
 
ba80154a1118
configuration option to control timing output for (co)datatypes
 
traytel 
parents: 
53126 
diff
changeset
 | 
97  | 
val bnf_timing: bool Config.T  | 
| 49435 | 98  | 
val user_policy: fact_policy -> Proof.context -> fact_policy  | 
| 52720 | 99  | 
val note_bnf_thms: fact_policy -> (binding -> binding) -> binding -> bnf -> Proof.context ->  | 
100  | 
Proof.context  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
101  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
102  | 
val print_bnfs: Proof.context -> unit  | 
| 
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: 
54490 
diff
changeset
 | 
103  | 
val prepare_def: const_policy -> (Proof.context -> fact_policy) -> (binding -> binding) ->  | 
| 
 
91a1e4aa7c80
command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
 
traytel 
parents: 
54490 
diff
changeset
 | 
104  | 
(Proof.context -> 'a -> typ) -> (Proof.context -> 'b -> term) -> typ list option ->  | 
| 
 
91a1e4aa7c80
command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
 
traytel 
parents: 
54490 
diff
changeset
 | 
105  | 
binding -> binding -> binding list ->  | 
| 
 
91a1e4aa7c80
command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
 
traytel 
parents: 
54490 
diff
changeset
 | 
106  | 
(((((binding * 'a) * 'b) * 'b list) * 'b) * 'b list) * 'b option -> Proof.context ->  | 
| 
 
91a1e4aa7c80
command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
 
traytel 
parents: 
54490 
diff
changeset
 | 
107  | 
string * term list *  | 
| 55197 | 108  | 
((Proof.context -> thm list -> tactic) option * term list list) *  | 
| 
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: 
54490 
diff
changeset
 | 
109  | 
((thm list -> thm list list) -> thm list list -> Proof.context -> bnf * local_theory) *  | 
| 
 
91a1e4aa7c80
command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
 
traytel 
parents: 
54490 
diff
changeset
 | 
110  | 
local_theory * thm list  | 
| 
 
91a1e4aa7c80
command for axiomatic declaration of bnfs---allow theoreticians to fix an arbitrary type with functorial structure and work abstractly on it
 
traytel 
parents: 
54490 
diff
changeset
 | 
111  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
112  | 
val define_bnf_consts: const_policy -> fact_policy -> typ list option ->  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
113  | 
binding -> binding -> binding list ->  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
114  | 
(((((binding * typ) * term) * term list) * term) * term list) * term option -> local_theory ->  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
115  | 
((typ list * typ list * typ list * typ) *  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
116  | 
(term * term list * term * (int list * term) list * term) *  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
117  | 
(thm * thm list * thm * thm list * thm) *  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
118  | 
((typ list -> typ list -> typ list -> term) *  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
119  | 
(typ list -> typ list -> term -> term) *  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
120  | 
(typ list -> typ list -> typ -> typ) *  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
121  | 
(typ list -> typ list -> typ list -> term) *  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
122  | 
(typ list -> typ list -> typ list -> term))) * local_theory  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
123  | 
|
| 49018 | 124  | 
val bnf_def: const_policy -> (Proof.context -> fact_policy) -> (binding -> binding) ->  | 
| 55197 | 125  | 
(Proof.context -> tactic) list ->  | 
126  | 
(Proof.context -> tactic) -> typ list option -> binding ->  | 
|
| 
51767
 
bbcdd8519253
honor user-specified name for relator + generalize syntax
 
blanchet 
parents: 
51766 
diff
changeset
 | 
127  | 
binding -> binding list ->  | 
| 54421 | 128  | 
(((((binding * typ) * term) * term list) * term) * term list) * term option ->  | 
| 
51837
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
129  | 
local_theory -> bnf * local_theory  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
130  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
131  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
132  | 
structure BNF_Def : BNF_DEF =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
133  | 
struct  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
134  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
135  | 
open BNF_Util  | 
| 49463 | 136  | 
open BNF_Tactics  | 
| 49284 | 137  | 
open BNF_Def_Tactics  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
138  | 
|
| 
54624
 
36301c99ed26
revert making 'map_cong' a 'cong' -- it breaks too many proofs in the AFP
 
blanchet 
parents: 
54620 
diff
changeset
 | 
139  | 
val fundefcong_attrs = @{attributes [fundef_cong]};
 | 
| 51765 | 140  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
141  | 
type axioms = {
 | 
| 53270 | 142  | 
map_id0: thm,  | 
| 53287 | 143  | 
map_comp0: thm,  | 
| 
51761
 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 
blanchet 
parents: 
51758 
diff
changeset
 | 
144  | 
map_cong0: thm,  | 
| 53289 | 145  | 
set_map0: thm list,  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
146  | 
bd_card_order: thm,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
147  | 
bd_cinfinite: thm,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
148  | 
set_bd: thm list,  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
149  | 
le_rel_OO: thm,  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
150  | 
rel_OO_Grp: thm  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
151  | 
};  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
152  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
153  | 
fun mk_axioms' ((((((((id, comp), cong), map), c_o), cinf), set_bd), le_rel_OO), rel) =  | 
| 53289 | 154  | 
  {map_id0 = id, map_comp0 = comp, map_cong0 = cong, set_map0 = map, bd_card_order = c_o,
 | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
155  | 
bd_cinfinite = cinf, set_bd = set_bd, le_rel_OO = le_rel_OO, rel_OO_Grp = rel};  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
156  | 
|
| 
51930
 
52fd62618631
prefer explicitly qualified exceptions, which is particular important for robust handlers;
 
wenzelm 
parents: 
51917 
diff
changeset
 | 
157  | 
fun dest_cons [] = raise List.Empty  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
158  | 
| dest_cons (x :: xs) = (x, xs);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
159  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
160  | 
fun mk_axioms n thms = thms  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
161  | 
|> map the_single  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
162  | 
|> dest_cons  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
163  | 
||>> dest_cons  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
164  | 
||>> dest_cons  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
165  | 
||>> chop n  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
166  | 
||>> dest_cons  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
167  | 
||>> dest_cons  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
168  | 
||>> chop n  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
169  | 
||>> dest_cons  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
170  | 
||> the_single  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
171  | 
|> mk_axioms';  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
172  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
173  | 
fun zip_axioms mid mcomp mcong smap bdco bdinf sbd le_rel_OO rel =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
174  | 
[mid, mcomp, mcong] @ smap @ [bdco, bdinf] @ sbd @ [le_rel_OO, rel];  | 
| 
49460
 
4dd451a075ce
fixed infinite loop with trivial rel_O_Gr + tuning
 
blanchet 
parents: 
49459 
diff
changeset
 | 
175  | 
|
| 53289 | 176  | 
fun dest_axioms {map_id0, map_comp0, map_cong0, set_map0, bd_card_order, bd_cinfinite, set_bd,
 | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
177  | 
le_rel_OO, rel_OO_Grp} =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
178  | 
zip_axioms map_id0 map_comp0 map_cong0 set_map0 bd_card_order bd_cinfinite set_bd le_rel_OO  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
179  | 
rel_OO_Grp;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
180  | 
|
| 53289 | 181  | 
fun map_axioms f {map_id0, map_comp0, map_cong0, set_map0, bd_card_order, bd_cinfinite, set_bd,
 | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
182  | 
le_rel_OO, rel_OO_Grp} =  | 
| 53270 | 183  | 
  {map_id0 = f map_id0,
 | 
| 53287 | 184  | 
map_comp0 = f map_comp0,  | 
| 
51761
 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 
blanchet 
parents: 
51758 
diff
changeset
 | 
185  | 
map_cong0 = f map_cong0,  | 
| 53289 | 186  | 
set_map0 = map f set_map0,  | 
| 49463 | 187  | 
bd_card_order = f bd_card_order,  | 
188  | 
bd_cinfinite = f bd_cinfinite,  | 
|
189  | 
set_bd = map f set_bd,  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
190  | 
le_rel_OO = f le_rel_OO,  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
191  | 
rel_OO_Grp = f rel_OO_Grp};  | 
| 
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  | 
val morph_axioms = map_axioms o Morphism.thm;  | 
| 
 
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  | 
type defs = {
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
196  | 
map_def: thm,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
197  | 
set_defs: thm list,  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
198  | 
rel_def: thm  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
199  | 
}  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
200  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
201  | 
fun mk_defs map sets rel = {map_def = map, set_defs = sets, rel_def = rel};
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
202  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
203  | 
fun map_defs f {map_def, set_defs, rel_def} =
 | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
204  | 
  {map_def = f map_def, set_defs = map f set_defs, rel_def = f rel_def};
 | 
| 
48975
 
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  | 
val morph_defs = map_defs o Morphism.thm;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
207  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
208  | 
type facts = {
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
209  | 
bd_Card_order: thm,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
210  | 
bd_Cinfinite: thm,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
211  | 
bd_Cnotzero: thm,  | 
| 
51766
 
f19a4d0ab1bf
renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
 
blanchet 
parents: 
51765 
diff
changeset
 | 
212  | 
collect_set_map: thm lazy,  | 
| 
52635
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
213  | 
in_bd: thm lazy,  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
214  | 
in_cong: thm lazy,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
215  | 
in_mono: thm lazy,  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
216  | 
in_rel: thm lazy,  | 
| 53288 | 217  | 
map_comp: thm lazy,  | 
| 51762 | 218  | 
map_cong: thm lazy,  | 
| 53285 | 219  | 
map_id: thm lazy,  | 
| 
52719
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
220  | 
map_transfer: thm lazy,  | 
| 
49591
 
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
 
blanchet 
parents: 
49585 
diff
changeset
 | 
221  | 
rel_eq: thm lazy,  | 
| 
49537
 
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
 
blanchet 
parents: 
49536 
diff
changeset
 | 
222  | 
rel_flip: thm lazy,  | 
| 53290 | 223  | 
set_map: thm lazy list,  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
224  | 
rel_cong: thm lazy,  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
225  | 
rel_mono: thm lazy,  | 
| 51916 | 226  | 
rel_mono_strong: thm lazy,  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
227  | 
rel_Grp: thm lazy,  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
228  | 
rel_conversep: thm lazy,  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
229  | 
rel_OO: thm lazy  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
230  | 
};  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
231  | 
|
| 
52635
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
232  | 
fun mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong in_mono in_rel  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
233  | 
map_comp map_cong map_id map_transfer rel_eq rel_flip set_map rel_cong rel_mono  | 
| 
52719
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
234  | 
    rel_mono_strong rel_Grp rel_conversep rel_OO = {
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
235  | 
bd_Card_order = bd_Card_order,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
236  | 
bd_Cinfinite = bd_Cinfinite,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
237  | 
bd_Cnotzero = bd_Cnotzero,  | 
| 
51766
 
f19a4d0ab1bf
renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
 
blanchet 
parents: 
51765 
diff
changeset
 | 
238  | 
collect_set_map = collect_set_map,  | 
| 
52635
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
239  | 
in_bd = in_bd,  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
240  | 
in_cong = in_cong,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
241  | 
in_mono = in_mono,  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
242  | 
in_rel = in_rel,  | 
| 53288 | 243  | 
map_comp = map_comp,  | 
| 51762 | 244  | 
map_cong = map_cong,  | 
| 53285 | 245  | 
map_id = map_id,  | 
| 
52719
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
246  | 
map_transfer = map_transfer,  | 
| 
49591
 
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
 
blanchet 
parents: 
49585 
diff
changeset
 | 
247  | 
rel_eq = rel_eq,  | 
| 
49537
 
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
 
blanchet 
parents: 
49536 
diff
changeset
 | 
248  | 
rel_flip = rel_flip,  | 
| 53290 | 249  | 
set_map = set_map,  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
250  | 
rel_cong = rel_cong,  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
251  | 
rel_mono = rel_mono,  | 
| 51916 | 252  | 
rel_mono_strong = rel_mono_strong,  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
253  | 
rel_Grp = rel_Grp,  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
254  | 
rel_conversep = rel_conversep,  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
255  | 
rel_OO = rel_OO};  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
256  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
257  | 
fun map_facts f {
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
258  | 
bd_Card_order,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
259  | 
bd_Cinfinite,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
260  | 
bd_Cnotzero,  | 
| 
51766
 
f19a4d0ab1bf
renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
 
blanchet 
parents: 
51765 
diff
changeset
 | 
261  | 
collect_set_map,  | 
| 
52635
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
262  | 
in_bd,  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
263  | 
in_cong,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
264  | 
in_mono,  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
265  | 
in_rel,  | 
| 53288 | 266  | 
map_comp,  | 
| 51762 | 267  | 
map_cong,  | 
| 53285 | 268  | 
map_id,  | 
| 
52719
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
269  | 
map_transfer,  | 
| 
49591
 
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
 
blanchet 
parents: 
49585 
diff
changeset
 | 
270  | 
rel_eq,  | 
| 
49537
 
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
 
blanchet 
parents: 
49536 
diff
changeset
 | 
271  | 
rel_flip,  | 
| 53290 | 272  | 
set_map,  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
273  | 
rel_cong,  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
274  | 
rel_mono,  | 
| 51916 | 275  | 
rel_mono_strong,  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
276  | 
rel_Grp,  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
277  | 
rel_conversep,  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
278  | 
rel_OO} =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
279  | 
  {bd_Card_order = f bd_Card_order,
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
280  | 
bd_Cinfinite = f bd_Cinfinite,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
281  | 
bd_Cnotzero = f bd_Cnotzero,  | 
| 
51766
 
f19a4d0ab1bf
renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
 
blanchet 
parents: 
51765 
diff
changeset
 | 
282  | 
collect_set_map = Lazy.map f collect_set_map,  | 
| 
52635
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
283  | 
in_bd = Lazy.map f in_bd,  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
284  | 
in_cong = Lazy.map f in_cong,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
285  | 
in_mono = Lazy.map f in_mono,  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
286  | 
in_rel = Lazy.map f in_rel,  | 
| 53288 | 287  | 
map_comp = Lazy.map f map_comp,  | 
| 51762 | 288  | 
map_cong = Lazy.map f map_cong,  | 
| 53285 | 289  | 
map_id = Lazy.map f map_id,  | 
| 
52719
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
290  | 
map_transfer = Lazy.map f map_transfer,  | 
| 
49591
 
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
 
blanchet 
parents: 
49585 
diff
changeset
 | 
291  | 
rel_eq = Lazy.map f rel_eq,  | 
| 
49537
 
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
 
blanchet 
parents: 
49536 
diff
changeset
 | 
292  | 
rel_flip = Lazy.map f rel_flip,  | 
| 53290 | 293  | 
set_map = map (Lazy.map f) set_map,  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
294  | 
rel_cong = Lazy.map f rel_cong,  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
295  | 
rel_mono = Lazy.map f rel_mono,  | 
| 51916 | 296  | 
rel_mono_strong = Lazy.map f rel_mono_strong,  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
297  | 
rel_Grp = Lazy.map f rel_Grp,  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
298  | 
rel_conversep = Lazy.map f rel_conversep,  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
299  | 
rel_OO = Lazy.map f rel_OO};  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
300  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
301  | 
val morph_facts = map_facts o Morphism.thm;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
302  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
303  | 
type nonemptiness_witness = {
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
304  | 
I: int list,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
305  | 
wit: term,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
306  | 
prop: thm list  | 
| 
 
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  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
309  | 
fun mk_witness (I, wit) prop = {I = I, wit = wit, prop = prop};
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
310  | 
fun map_witness f g {I, wit, prop} = {I = I, wit = f wit, prop = map g prop};
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
311  | 
fun morph_witness phi = map_witness (Morphism.term phi) (Morphism.thm phi);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
312  | 
|
| 
51837
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
313  | 
datatype bnf = BNF of {
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
314  | 
name: binding,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
315  | 
T: typ,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
316  | 
live: int,  | 
| 53261 | 317  | 
lives: typ list, (*source type variables of map*)  | 
318  | 
lives': typ list, (*target type variables of map*)  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
319  | 
dead: int,  | 
| 53261 | 320  | 
deads: typ list,  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
321  | 
map: term,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
322  | 
sets: term list,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
323  | 
bd: term,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
324  | 
axioms: axioms,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
325  | 
defs: defs,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
326  | 
facts: facts,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
327  | 
nwits: int,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
328  | 
wits: nonemptiness_witness list,  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
329  | 
rel: term  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
330  | 
};  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
331  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
332  | 
(* getters *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
333  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
334  | 
fun rep_bnf (BNF bnf) = bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
335  | 
val name_of_bnf = #name o rep_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
336  | 
val T_of_bnf = #T o rep_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
337  | 
fun mk_T_of_bnf Ds Ts bnf =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
338  | 
let val bnf_rep = rep_bnf bnf  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
339  | 
in Term.typ_subst_atomic ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#T bnf_rep) end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
340  | 
val live_of_bnf = #live o rep_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
341  | 
val lives_of_bnf = #lives o rep_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
342  | 
val dead_of_bnf = #dead o rep_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
343  | 
val deads_of_bnf = #deads o rep_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
344  | 
val axioms_of_bnf = #axioms o rep_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
345  | 
val facts_of_bnf = #facts o rep_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
346  | 
val nwits_of_bnf = #nwits o rep_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
347  | 
val wits_of_bnf = #wits o rep_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
348  | 
|
| 53031 | 349  | 
fun flatten_type_args_of_bnf bnf dead_x xs =  | 
350  | 
let  | 
|
351  | 
val Type (_, Ts) = T_of_bnf bnf;  | 
|
352  | 
val lives = lives_of_bnf bnf;  | 
|
353  | 
val deads = deads_of_bnf bnf;  | 
|
354  | 
in  | 
|
| 
55480
 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 
blanchet 
parents: 
55444 
diff
changeset
 | 
355  | 
permute_like_unique (op =) (deads @ lives) Ts (replicate (length deads) dead_x @ xs)  | 
| 53031 | 356  | 
end;  | 
357  | 
||
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
358  | 
(*terms*)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
359  | 
val map_of_bnf = #map o rep_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
360  | 
val sets_of_bnf = #sets o rep_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
361  | 
fun mk_map_of_bnf Ds Ts Us bnf =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
362  | 
let val bnf_rep = rep_bnf bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
363  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
364  | 
Term.subst_atomic_types  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
365  | 
((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#map bnf_rep)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
366  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
367  | 
fun mk_sets_of_bnf Dss Tss bnf =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
368  | 
let val bnf_rep = rep_bnf bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
369  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
370  | 
map2 (fn (Ds, Ts) => Term.subst_atomic_types  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
371  | 
((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts))) (Dss ~~ Tss) (#sets bnf_rep)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
372  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
373  | 
val bd_of_bnf = #bd o rep_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
374  | 
fun mk_bd_of_bnf Ds Ts bnf =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
375  | 
let val bnf_rep = rep_bnf bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
376  | 
in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#bd bnf_rep) end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
377  | 
fun mk_wits_of_bnf Dss Tss bnf =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
378  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
379  | 
val bnf_rep = rep_bnf bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
380  | 
val wits = map (fn x => (#I x, #wit x)) (#wits bnf_rep);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
381  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
382  | 
map2 (fn (Ds, Ts) => apsnd (Term.subst_atomic_types  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
383  | 
((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)))) (Dss ~~ Tss) wits  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
384  | 
end;  | 
| 49507 | 385  | 
val rel_of_bnf = #rel o rep_bnf;  | 
386  | 
fun mk_rel_of_bnf Ds Ts Us bnf =  | 
|
| 49462 | 387  | 
let val bnf_rep = rep_bnf bnf;  | 
388  | 
in  | 
|
389  | 
Term.subst_atomic_types  | 
|
| 49507 | 390  | 
((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#rel bnf_rep)  | 
| 49462 | 391  | 
end;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
392  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
393  | 
(*thms*)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
394  | 
val bd_card_order_of_bnf = #bd_card_order o #axioms o rep_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
395  | 
val bd_cinfinite_of_bnf = #bd_cinfinite o #axioms o rep_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
396  | 
val bd_Card_order_of_bnf = #bd_Card_order o #facts o rep_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
397  | 
val bd_Cinfinite_of_bnf = #bd_Cinfinite o #facts o rep_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
398  | 
val bd_Cnotzero_of_bnf = #bd_Cnotzero o #facts o rep_bnf;  | 
| 
51766
 
f19a4d0ab1bf
renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
 
blanchet 
parents: 
51765 
diff
changeset
 | 
399  | 
val collect_set_map_of_bnf = Lazy.force o #collect_set_map o #facts o rep_bnf;  | 
| 
52635
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
400  | 
val in_bd_of_bnf = Lazy.force o #in_bd o #facts o rep_bnf;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
401  | 
val in_cong_of_bnf = Lazy.force o #in_cong o #facts o rep_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
402  | 
val in_mono_of_bnf = Lazy.force o #in_mono o #facts o rep_bnf;  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
403  | 
val in_rel_of_bnf = Lazy.force o #in_rel o #facts o rep_bnf;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
404  | 
val map_def_of_bnf = #map_def o #defs o rep_bnf;  | 
| 53270 | 405  | 
val map_id0_of_bnf = #map_id0 o #axioms o rep_bnf;  | 
| 53285 | 406  | 
val map_id_of_bnf = Lazy.force o #map_id o #facts o rep_bnf;  | 
| 53287 | 407  | 
val map_comp0_of_bnf = #map_comp0 o #axioms o rep_bnf;  | 
| 53288 | 408  | 
val map_comp_of_bnf = Lazy.force o #map_comp o #facts o rep_bnf;  | 
| 
51761
 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 
blanchet 
parents: 
51758 
diff
changeset
 | 
409  | 
val map_cong0_of_bnf = #map_cong0 o #axioms o rep_bnf;  | 
| 51762 | 410  | 
val map_cong_of_bnf = Lazy.force o #map_cong o #facts o rep_bnf;  | 
| 52731 | 411  | 
val map_transfer_of_bnf = Lazy.force o #map_transfer o #facts o rep_bnf;  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
412  | 
val le_rel_OO_of_bnf = #le_rel_OO o #axioms o rep_bnf;  | 
| 49507 | 413  | 
val rel_def_of_bnf = #rel_def o #defs o rep_bnf;  | 
| 
49591
 
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
 
blanchet 
parents: 
49585 
diff
changeset
 | 
414  | 
val rel_eq_of_bnf = Lazy.force o #rel_eq o #facts o rep_bnf;  | 
| 
49537
 
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
 
blanchet 
parents: 
49536 
diff
changeset
 | 
415  | 
val rel_flip_of_bnf = Lazy.force o #rel_flip o #facts o rep_bnf;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
416  | 
val set_bd_of_bnf = #set_bd o #axioms o rep_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
417  | 
val set_defs_of_bnf = #set_defs o #defs o rep_bnf;  | 
| 53289 | 418  | 
val set_map0_of_bnf = #set_map0 o #axioms o rep_bnf;  | 
| 53290 | 419  | 
val set_map_of_bnf = map Lazy.force o #set_map o #facts o rep_bnf;  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
420  | 
val rel_cong_of_bnf = Lazy.force o #rel_cong o #facts o rep_bnf;  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
421  | 
val rel_mono_of_bnf = Lazy.force o #rel_mono o #facts o rep_bnf;  | 
| 51916 | 422  | 
val rel_mono_strong_of_bnf = Lazy.force o #rel_mono_strong o #facts o rep_bnf;  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
423  | 
val rel_Grp_of_bnf = Lazy.force o #rel_Grp o #facts o rep_bnf;  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
424  | 
val rel_conversep_of_bnf = Lazy.force o #rel_conversep o #facts o rep_bnf;  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
425  | 
val rel_OO_of_bnf = Lazy.force o #rel_OO o #facts o rep_bnf;  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
426  | 
val rel_OO_Grp_of_bnf = #rel_OO_Grp o #axioms o rep_bnf;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
427  | 
val wit_thms_of_bnf = maps #prop o wits_of_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
428  | 
val wit_thmss_of_bnf = map #prop o wits_of_bnf;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
429  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
430  | 
fun mk_bnf name T live lives lives' dead deads map sets bd axioms defs facts wits rel =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
431  | 
  BNF {name = name, T = T,
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
432  | 
live = live, lives = lives, lives' = lives', dead = dead, deads = deads,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
433  | 
map = map, sets = sets, bd = bd,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
434  | 
axioms = axioms, defs = defs, facts = facts,  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
435  | 
nwits = length wits, wits = wits, rel = rel};  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
436  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
437  | 
fun morph_bnf phi (BNF {name = name, T = T, live = live, lives = lives, lives' = lives',
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
438  | 
dead = dead, deads = deads, map = map, sets = sets, bd = bd,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
439  | 
axioms = axioms, defs = defs, facts = facts,  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
440  | 
nwits = nwits, wits = wits, rel = rel}) =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
441  | 
  BNF {name = Morphism.binding phi name, T = Morphism.typ phi T,
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
442  | 
live = live, lives = List.map (Morphism.typ phi) lives,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
443  | 
lives' = List.map (Morphism.typ phi) lives',  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
444  | 
dead = dead, deads = List.map (Morphism.typ phi) deads,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
445  | 
map = Morphism.term phi map, sets = List.map (Morphism.term phi) sets,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
446  | 
bd = Morphism.term phi bd,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
447  | 
axioms = morph_axioms phi axioms,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
448  | 
defs = morph_defs phi defs,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
449  | 
facts = morph_facts phi facts,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
450  | 
nwits = nwits,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
451  | 
wits = List.map (morph_witness phi) wits,  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
452  | 
rel = Morphism.term phi rel};  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
453  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
454  | 
structure Data = Generic_Data  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
455  | 
(  | 
| 
51837
 
087498724486
lowercase type constructor, for consistency (cf. fp_result not FP_result nor FP_Result)
 
blanchet 
parents: 
51836 
diff
changeset
 | 
456  | 
type T = bnf Symtab.table;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
457  | 
val empty = Symtab.empty;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
458  | 
val extend = I;  | 
| 
55394
 
d5ebe055b599
more liberal merging of BNFs and constructor sugar
 
blanchet 
parents: 
55356 
diff
changeset
 | 
459  | 
fun merge data : T = Symtab.merge (K true) data;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
460  | 
);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
461  | 
|
| 
53126
 
f4d2c64c7aa8
transfer stored fp_sugar theorems into the "current" theory when retrieving an fp_sugar (avoids non-trivial merges)
 
traytel 
parents: 
53040 
diff
changeset
 | 
462  | 
fun bnf_of ctxt =  | 
| 
 
f4d2c64c7aa8
transfer stored fp_sugar theorems into the "current" theory when retrieving an fp_sugar (avoids non-trivial merges)
 
traytel 
parents: 
53040 
diff
changeset
 | 
463  | 
Symtab.lookup (Data.get (Context.Proof ctxt))  | 
| 54740 | 464  | 
#> Option.map (morph_bnf (Morphism.transfer_morphism (Proof_Context.theory_of ctxt)));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
465  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
466  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
467  | 
(* Utilities *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
468  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
469  | 
fun normalize_set insts instA set =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
470  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
471  | 
val (T, T') = dest_funT (fastype_of set);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
472  | 
val A = fst (Term.dest_TVar (HOLogic.dest_setT T'));  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
473  | 
val params = Term.add_tvar_namesT T [];  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
474  | 
in Term.subst_TVars ((A :: params) ~~ (instA :: insts)) set end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
475  | 
|
| 49507 | 476  | 
fun normalize_rel ctxt instTs instA instB rel =  | 
| 49462 | 477  | 
let  | 
478  | 
val thy = Proof_Context.theory_of ctxt;  | 
|
479  | 
val tyenv =  | 
|
| 49507 | 480  | 
Sign.typ_match thy (fastype_of rel, Library.foldr (op -->) (instTs, mk_pred2T instA instB))  | 
| 49463 | 481  | 
Vartab.empty;  | 
| 49507 | 482  | 
in Envir.subst_term (tyenv, Vartab.empty) rel end  | 
| 49453 | 483  | 
handle Type.TYPE_MATCH => error "Bad relator";  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
484  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
485  | 
fun normalize_wit insts CA As wit =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
486  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
487  | 
    fun strip_param (Ts, T as Type (@{type_name fun}, [T1, T2])) =
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
488  | 
if Type.raw_instance (CA, T) then (Ts, T) else strip_param (T1 :: Ts, T2)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
489  | 
| strip_param x = x;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
490  | 
val (Ts, T) = strip_param ([], fastype_of wit);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
491  | 
val subst = Term.add_tvar_namesT T [] ~~ insts;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
492  | 
fun find y = find_index (fn x => x = y) As;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
493  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
494  | 
(map (find o Term.typ_subst_TVars subst) (rev Ts), Term.subst_TVars subst wit)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
495  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
496  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
497  | 
fun minimize_wits wits =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
498  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
499  | 
fun minimize done [] = done  | 
| 49103 | 500  | 
| minimize done ((I, wit) :: todo) =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
501  | 
if exists (fn (J, _) => subset (op =) (J, I)) (done @ todo)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
502  | 
then minimize done todo  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
503  | 
else minimize ((I, wit) :: done) todo;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
504  | 
in minimize [] wits end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
505  | 
|
| 54236 | 506  | 
fun mk_map live Ts Us t =  | 
507  | 
let val (Type (_, Ts0), Type (_, Us0)) = strip_typeN (live + 1) (fastype_of t) |>> List.last in  | 
|
508  | 
Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t  | 
|
509  | 
end;  | 
|
510  | 
||
511  | 
fun mk_rel live Ts Us t =  | 
|
512  | 
let val [Type (_, Ts0), Type (_, Us0)] = binder_types (snd (strip_typeN live (fastype_of t))) in  | 
|
513  | 
Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t  | 
|
514  | 
end;  | 
|
515  | 
||
| 
54237
 
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
 
blanchet 
parents: 
54236 
diff
changeset
 | 
516  | 
fun build_map_or_rel mk const of_bnf dest ctxt build_simple =  | 
| 54236 | 517  | 
let  | 
518  | 
fun build (TU as (T, U)) =  | 
|
519  | 
if T = U then  | 
|
520  | 
const T  | 
|
521  | 
else  | 
|
522  | 
(case TU of  | 
|
523  | 
(Type (s, Ts), Type (s', Us)) =>  | 
|
524  | 
if s = s' then  | 
|
525  | 
let  | 
|
| 
54237
 
7cc6e286fe28
made sugared 'coinduct' theorem construction n2m-proof
 
blanchet 
parents: 
54236 
diff
changeset
 | 
526  | 
val bnf = the (bnf_of ctxt s);  | 
| 54236 | 527  | 
val live = live_of_bnf bnf;  | 
528  | 
val mapx = mk live Ts Us (of_bnf bnf);  | 
|
529  | 
val TUs' = map dest (fst (strip_typeN live (fastype_of mapx)));  | 
|
530  | 
in Term.list_comb (mapx, map build TUs') end  | 
|
531  | 
else  | 
|
532  | 
build_simple TU  | 
|
533  | 
| _ => build_simple TU);  | 
|
534  | 
in build end;  | 
|
535  | 
||
536  | 
val build_map = build_map_or_rel mk_map HOLogic.id_const map_of_bnf dest_funT;  | 
|
537  | 
val build_rel = build_map_or_rel mk_rel HOLogic.eq_const rel_of_bnf dest_pred2T;  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
538  | 
|
| 
54246
 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 
blanchet 
parents: 
54237 
diff
changeset
 | 
539  | 
fun map_flattened_map_args ctxt s map_args fs =  | 
| 
 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 
blanchet 
parents: 
54237 
diff
changeset
 | 
540  | 
let  | 
| 
 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 
blanchet 
parents: 
54237 
diff
changeset
 | 
541  | 
val flat_fs = flatten_type_args_of_bnf (the (bnf_of ctxt s)) Term.dummy fs;  | 
| 
 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 
blanchet 
parents: 
54237 
diff
changeset
 | 
542  | 
val flat_fs' = map_args flat_fs;  | 
| 
 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 
blanchet 
parents: 
54237 
diff
changeset
 | 
543  | 
in  | 
| 
55480
 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 
blanchet 
parents: 
55444 
diff
changeset
 | 
544  | 
permute_like_unique (op aconv) flat_fs fs flat_fs'  | 
| 
54246
 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 
blanchet 
parents: 
54237 
diff
changeset
 | 
545  | 
end;  | 
| 
 
8fdb4dc08ed1
split 'primrec_new' and 'primcorec' code (to ease bootstrapping, e.g. dependency on datatype 'String' in 'primcorec')
 
blanchet 
parents: 
54237 
diff
changeset
 | 
546  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
547  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
548  | 
(* Names *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
549  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
550  | 
val mapN = "map";  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
551  | 
val setN = "set";  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
552  | 
fun mk_setN i = setN ^ nonzero_string_of_int i;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
553  | 
val bdN = "bd";  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
554  | 
val witN = "wit";  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
555  | 
fun mk_witN i = witN ^ nonzero_string_of_int i;  | 
| 49507 | 556  | 
val relN = "rel";  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
557  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
558  | 
val bd_card_orderN = "bd_card_order";  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
559  | 
val bd_cinfiniteN = "bd_cinfinite";  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
560  | 
val bd_Card_orderN = "bd_Card_order";  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
561  | 
val bd_CinfiniteN = "bd_Cinfinite";  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
562  | 
val bd_CnotzeroN = "bd_Cnotzero";  | 
| 
51766
 
f19a4d0ab1bf
renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
 
blanchet 
parents: 
51765 
diff
changeset
 | 
563  | 
val collect_set_mapN = "collect_set_map";  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
564  | 
val in_bdN = "in_bd";  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
565  | 
val in_monoN = "in_mono";  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
566  | 
val in_relN = "in_rel";  | 
| 53270 | 567  | 
val map_id0N = "map_id0";  | 
| 53285 | 568  | 
val map_idN = "map_id";  | 
| 53287 | 569  | 
val map_comp0N = "map_comp0";  | 
| 53288 | 570  | 
val map_compN = "map_comp";  | 
| 
51761
 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 
blanchet 
parents: 
51758 
diff
changeset
 | 
571  | 
val map_cong0N = "map_cong0";  | 
| 51762 | 572  | 
val map_congN = "map_cong";  | 
| 
52719
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
573  | 
val map_transferN = "map_transfer";  | 
| 
49591
 
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
 
blanchet 
parents: 
49585 
diff
changeset
 | 
574  | 
val rel_eqN = "rel_eq";  | 
| 
49537
 
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
 
blanchet 
parents: 
49536 
diff
changeset
 | 
575  | 
val rel_flipN = "rel_flip";  | 
| 53289 | 576  | 
val set_map0N = "set_map0";  | 
| 53290 | 577  | 
val set_mapN = "set_map";  | 
| 
49537
 
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
 
blanchet 
parents: 
49536 
diff
changeset
 | 
578  | 
val set_bdN = "set_bd";  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
579  | 
val rel_GrpN = "rel_Grp";  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
580  | 
val rel_conversepN = "rel_conversep";  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
581  | 
val rel_monoN = "rel_mono"  | 
| 51916 | 582  | 
val rel_mono_strongN = "rel_mono_strong"  | 
| 54620 | 583  | 
val rel_comppN = "rel_compp";  | 
584  | 
val rel_compp_GrpN = "rel_compp_Grp";  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
585  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
586  | 
datatype const_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
587  | 
|
| 49538 | 588  | 
datatype fact_policy = Dont_Note | Note_Some | Note_All;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
589  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
590  | 
val bnf_note_all = Attrib.setup_config_bool @{binding bnf_note_all} (K false);
 | 
| 
53143
 
ba80154a1118
configuration option to control timing output for (co)datatypes
 
traytel 
parents: 
53126 
diff
changeset
 | 
591  | 
val bnf_timing = Attrib.setup_config_bool @{binding bnf_timing} (K false);
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
592  | 
|
| 49538 | 593  | 
fun user_policy policy ctxt = if Config.get ctxt bnf_note_all then Note_All else policy;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
594  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
595  | 
val smart_max_inline_size = 25; (*FUDGE*)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
596  | 
|
| 
54045
 
369a4a14583a
keep the qualification of bindings when noting bnf theorems
 
traytel 
parents: 
53561 
diff
changeset
 | 
597  | 
fun note_bnf_thms fact_policy qualify' bnf_b bnf =  | 
| 52720 | 598  | 
let  | 
599  | 
val axioms = axioms_of_bnf bnf;  | 
|
600  | 
val facts = facts_of_bnf bnf;  | 
|
601  | 
val wits = wits_of_bnf bnf;  | 
|
| 
54045
 
369a4a14583a
keep the qualification of bindings when noting bnf theorems
 
traytel 
parents: 
53561 
diff
changeset
 | 
602  | 
val qualify =  | 
| 
 
369a4a14583a
keep the qualification of bindings when noting bnf theorems
 
traytel 
parents: 
53561 
diff
changeset
 | 
603  | 
let val (_, qs, _) = Binding.dest bnf_b;  | 
| 
 
369a4a14583a
keep the qualification of bindings when noting bnf theorems
 
traytel 
parents: 
53561 
diff
changeset
 | 
604  | 
in fold_rev (fn (s, mand) => Binding.qualify mand s) qs #> qualify' end;  | 
| 52720 | 605  | 
in  | 
606  | 
(if fact_policy = Note_All then  | 
|
607  | 
let  | 
|
608  | 
val witNs = if length wits = 1 then [witN] else map mk_witN (1 upto length wits);  | 
|
609  | 
val notes =  | 
|
610  | 
[(bd_card_orderN, [#bd_card_order axioms]),  | 
|
611  | 
(bd_cinfiniteN, [#bd_cinfinite axioms]),  | 
|
612  | 
(bd_Card_orderN, [#bd_Card_order facts]),  | 
|
613  | 
(bd_CinfiniteN, [#bd_Cinfinite facts]),  | 
|
614  | 
(bd_CnotzeroN, [#bd_Cnotzero facts]),  | 
|
615  | 
(collect_set_mapN, [Lazy.force (#collect_set_map facts)]),  | 
|
616  | 
(in_bdN, [Lazy.force (#in_bd facts)]),  | 
|
617  | 
(in_monoN, [Lazy.force (#in_mono facts)]),  | 
|
618  | 
(in_relN, [Lazy.force (#in_rel facts)]),  | 
|
| 53287 | 619  | 
(map_comp0N, [#map_comp0 axioms]),  | 
| 53270 | 620  | 
(map_id0N, [#map_id0 axioms]),  | 
| 52720 | 621  | 
(map_transferN, [Lazy.force (#map_transfer facts)]),  | 
622  | 
(rel_mono_strongN, [Lazy.force (#rel_mono_strong facts)]),  | 
|
| 53289 | 623  | 
(set_map0N, #set_map0 axioms),  | 
| 52720 | 624  | 
(set_bdN, #set_bd axioms)] @  | 
625  | 
(witNs ~~ wit_thmss_of_bnf bnf)  | 
|
626  | 
|> map (fn (thmN, thms) =>  | 
|
| 53265 | 627  | 
((qualify (Binding.qualify true (Binding.name_of bnf_b) (Binding.name thmN)), []),  | 
| 52720 | 628  | 
[(thms, [])]));  | 
629  | 
in  | 
|
630  | 
Local_Theory.notes notes #> snd  | 
|
631  | 
end  | 
|
632  | 
else  | 
|
633  | 
I)  | 
|
634  | 
#> (if fact_policy <> Dont_Note then  | 
|
635  | 
let  | 
|
636  | 
val notes =  | 
|
| 53288 | 637  | 
[(map_compN, [Lazy.force (#map_comp facts)], []),  | 
| 52720 | 638  | 
(map_cong0N, [#map_cong0 axioms], []),  | 
| 
54624
 
36301c99ed26
revert making 'map_cong' a 'cong' -- it breaks too many proofs in the AFP
 
blanchet 
parents: 
54620 
diff
changeset
 | 
639  | 
(map_congN, [Lazy.force (#map_cong facts)], fundefcong_attrs),  | 
| 53285 | 640  | 
(map_idN, [Lazy.force (#map_id facts)], []),  | 
| 54620 | 641  | 
(rel_comppN, [Lazy.force (#rel_OO facts)], []),  | 
642  | 
(rel_compp_GrpN, no_refl [#rel_OO_Grp axioms], []),  | 
|
643  | 
(rel_conversepN, [Lazy.force (#rel_conversep facts)], []),  | 
|
| 52720 | 644  | 
(rel_eqN, [Lazy.force (#rel_eq facts)], []),  | 
645  | 
(rel_flipN, [Lazy.force (#rel_flip facts)], []),  | 
|
646  | 
(rel_GrpN, [Lazy.force (#rel_Grp facts)], []),  | 
|
647  | 
(rel_monoN, [Lazy.force (#rel_mono facts)], []),  | 
|
| 54620 | 648  | 
(set_mapN, map Lazy.force (#set_map facts), [])]  | 
| 52720 | 649  | 
|> filter_out (null o #2)  | 
650  | 
|> map (fn (thmN, thms, attrs) =>  | 
|
| 53265 | 651  | 
((qualify (Binding.qualify true (Binding.name_of bnf_b) (Binding.name thmN)),  | 
| 52720 | 652  | 
attrs), [(thms, [])]));  | 
653  | 
in  | 
|
654  | 
Local_Theory.notes notes #> snd  | 
|
655  | 
end  | 
|
656  | 
else  | 
|
657  | 
I)  | 
|
658  | 
end;  | 
|
659  | 
||
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
660  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
661  | 
(* Define new BNFs *)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
662  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
663  | 
fun define_bnf_consts const_policy fact_policy Ds_opt map_b rel_b set_bs  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
664  | 
((((((bnf_b, T_rhs), map_rhs), set_rhss), bd_rhs), wit_rhss), rel_rhs_opt) no_defs_lthy =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
665  | 
let  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
666  | 
val live = length set_rhss;  | 
| 53265 | 667  | 
|
| 54156 | 668  | 
val def_qualify = Binding.conceal o Binding.qualify false (Binding.name_of bnf_b);  | 
| 53265 | 669  | 
|
| 
54490
 
930409d43211
use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
 
blanchet 
parents: 
54426 
diff
changeset
 | 
670  | 
fun mk_prefix_binding pre = Binding.prefix_name (pre ^ "_") bnf_b;  | 
| 
49434
 
433dc7e028c8
separated registration of BNFs from bnf_def (BNFs are now stored only for bnf_def and (co)data commands)
 
traytel 
parents: 
49430 
diff
changeset
 | 
671  | 
|
| 49463 | 672  | 
fun maybe_define user_specified (b, rhs) lthy =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
673  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
674  | 
val inline =  | 
| 49538 | 675  | 
(user_specified orelse fact_policy = Dont_Note) andalso  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
676  | 
(case const_policy of  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
677  | 
Dont_Inline => false  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
678  | 
| Hardly_Inline => Term.is_Free rhs orelse Term.is_Const rhs  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
679  | 
| Smart_Inline => Term.size_of_term rhs <= smart_max_inline_size  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
680  | 
| Do_Inline => true)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
681  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
682  | 
if inline then  | 
| 
49460
 
4dd451a075ce
fixed infinite loop with trivial rel_O_Gr + tuning
 
blanchet 
parents: 
49459 
diff
changeset
 | 
683  | 
((rhs, Drule.reflexive_thm), lthy)  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
684  | 
else  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
685  | 
let val b = b () in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
686  | 
apfst (apsnd snd) (Local_Theory.define ((b, NoSyn), ((Thm.def_binding b, []), rhs))  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
687  | 
lthy)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
688  | 
end  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
689  | 
end;  | 
| 49459 | 690  | 
|
691  | 
fun maybe_restore lthy_old lthy =  | 
|
692  | 
lthy |> not (pointer_eq (lthy_old, lthy)) ? Local_Theory.restore;  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
693  | 
|
| 51758 | 694  | 
val map_bind_def =  | 
| 
54490
 
930409d43211
use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
 
blanchet 
parents: 
54426 
diff
changeset
 | 
695  | 
(fn () => def_qualify (if Binding.is_empty map_b then mk_prefix_binding mapN else map_b),  | 
| 53265 | 696  | 
map_rhs);  | 
| 49459 | 697  | 
val set_binds_defs =  | 
698  | 
let  | 
|
| 51757 | 699  | 
fun set_name i get_b =  | 
700  | 
(case try (nth set_bs) (i - 1) of  | 
|
701  | 
SOME b => if Binding.is_empty b then get_b else K b  | 
|
| 53265 | 702  | 
| NONE => get_b) #> def_qualify;  | 
| 
54490
 
930409d43211
use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
 
blanchet 
parents: 
54426 
diff
changeset
 | 
703  | 
val bs = if live = 1 then [set_name 1 (fn () => mk_prefix_binding setN)]  | 
| 
 
930409d43211
use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
 
blanchet 
parents: 
54426 
diff
changeset
 | 
704  | 
else map (fn i => set_name i (fn () => mk_prefix_binding (mk_setN i))) (1 upto live);  | 
| 51757 | 705  | 
in bs ~~ set_rhss end;  | 
| 
54490
 
930409d43211
use suffix '_list' etc. instead of prefix 'list_' for constants not present in the old package
 
blanchet 
parents: 
54426 
diff
changeset
 | 
706  | 
val bd_bind_def = (fn () => def_qualify (mk_prefix_binding bdN), bd_rhs);  | 
| 49459 | 707  | 
|
| 
54189
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
708  | 
val ((((bnf_map_term, raw_map_def),  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
709  | 
(bnf_set_terms, raw_set_defs)),  | 
| 
54189
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
710  | 
(bnf_bd_term, raw_bd_def)), (lthy, lthy_old)) =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
711  | 
no_defs_lthy  | 
| 49463 | 712  | 
|> maybe_define true map_bind_def  | 
713  | 
||>> apfst split_list o fold_map (maybe_define true) set_binds_defs  | 
|
714  | 
||>> maybe_define true bd_bind_def  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
715  | 
||> `(maybe_restore no_defs_lthy);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
716  | 
|
| 49459 | 717  | 
val phi = Proof_Context.export_morphism lthy_old lthy;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
718  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
719  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
720  | 
val bnf_map_def = Morphism.thm phi raw_map_def;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
721  | 
val bnf_set_defs = map (Morphism.thm phi) raw_set_defs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
722  | 
val bnf_bd_def = Morphism.thm phi raw_bd_def;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
723  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
724  | 
val bnf_map = Morphism.term phi bnf_map_term;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
725  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
726  | 
(*TODO: handle errors*)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
727  | 
(*simple shape analysis of a map function*)  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
728  | 
val ((alphas, betas), (Calpha, _)) =  | 
| 
49395
 
323414474c1f
use strip_typeN in bnf_def (instead of repairing strip_type)
 
traytel 
parents: 
49386 
diff
changeset
 | 
729  | 
fastype_of bnf_map  | 
| 
 
323414474c1f
use strip_typeN in bnf_def (instead of repairing strip_type)
 
traytel 
parents: 
49386 
diff
changeset
 | 
730  | 
|> strip_typeN live  | 
| 
 
323414474c1f
use strip_typeN in bnf_def (instead of repairing strip_type)
 
traytel 
parents: 
49386 
diff
changeset
 | 
731  | 
|>> map_split dest_funT  | 
| 
 
323414474c1f
use strip_typeN in bnf_def (instead of repairing strip_type)
 
traytel 
parents: 
49386 
diff
changeset
 | 
732  | 
||> dest_funT  | 
| 
 
323414474c1f
use strip_typeN in bnf_def (instead of repairing strip_type)
 
traytel 
parents: 
49386 
diff
changeset
 | 
733  | 
handle TYPE _ => error "Bad map function";  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
734  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
735  | 
val Calpha_params = map TVar (Term.add_tvarsT Calpha []);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
736  | 
|
| 
54426
 
edb87aac9cca
prohibit locally fixed type variables in bnf definitions
 
traytel 
parents: 
54425 
diff
changeset
 | 
737  | 
val bnf_T = Morphism.typ phi T_rhs;  | 
| 
 
edb87aac9cca
prohibit locally fixed type variables in bnf definitions
 
traytel 
parents: 
54425 
diff
changeset
 | 
738  | 
val bad_args = Term.add_tfreesT bnf_T [];  | 
| 
 
edb87aac9cca
prohibit locally fixed type variables in bnf definitions
 
traytel 
parents: 
54425 
diff
changeset
 | 
739  | 
    val _ = if null bad_args then () else error ("Locally fixed type arguments " ^
 | 
| 
 
edb87aac9cca
prohibit locally fixed type variables in bnf definitions
 
traytel 
parents: 
54425 
diff
changeset
 | 
740  | 
commas_quote (map (Syntax.string_of_typ no_defs_lthy o TFree) bad_args));  | 
| 
 
edb87aac9cca
prohibit locally fixed type variables in bnf definitions
 
traytel 
parents: 
54425 
diff
changeset
 | 
741  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
742  | 
val bnf_sets =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
743  | 
map2 (normalize_set Calpha_params) alphas (map (Morphism.term phi) bnf_set_terms);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
744  | 
val bnf_bd =  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
745  | 
Term.subst_TVars (Term.add_tvar_namesT bnf_T [] ~~ Calpha_params)  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
746  | 
(Morphism.term phi bnf_bd_term);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
747  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
748  | 
(*TODO: assert Ds = (TVars of bnf_map) \ (alphas @ betas) as sets*)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
749  | 
val deads = (case Ds_opt of  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
750  | 
NONE => subtract (op =) (alphas @ betas) (map TVar (Term.add_tvars bnf_map []))  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
751  | 
| SOME Ds => map (Morphism.typ phi) Ds);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
752  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
753  | 
(*TODO: further checks of type of bnf_map*)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
754  | 
(*TODO: check types of bnf_sets*)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
755  | 
(*TODO: check type of bnf_bd*)  | 
| 49507 | 756  | 
(*TODO: check type of bnf_rel*)  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
757  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
758  | 
fun mk_bnf_map Ds As' Bs' =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
759  | 
Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As') @ (betas ~~ Bs')) bnf_map;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
760  | 
fun mk_bnf_t Ds As' = Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As'));  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
761  | 
fun mk_bnf_T Ds As' = Term.typ_subst_atomic ((deads ~~ Ds) @ (alphas ~~ As'));  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
762  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
763  | 
val (((As, Bs), Ds), names_lthy) = lthy  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
764  | 
|> mk_TFrees live  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
765  | 
||>> mk_TFrees live  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
766  | 
||>> mk_TFrees (length deads);  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
767  | 
val RTs = map2 (curry HOLogic.mk_prodT) As Bs;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
768  | 
val pred2RTs = map2 mk_pred2T As Bs;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
769  | 
val (Rs, Rs') = names_lthy |> mk_Frees' "R" pred2RTs |> fst  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
770  | 
val CA = mk_bnf_T Ds As Calpha;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
771  | 
val CR = mk_bnf_T Ds RTs Calpha;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
772  | 
val setRs =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
773  | 
map3 (fn R => fn T => fn U =>  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
774  | 
HOLogic.Collect_const (HOLogic.mk_prodT (T, U)) $ HOLogic.mk_split R) Rs As Bs;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
775  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
776  | 
(*Grp (in (Collect (split R1) .. Collect (split Rn))) (map fst .. fst)^--1 OO  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
777  | 
Grp (in (Collect (split R1) .. Collect (split Rn))) (map snd .. snd)*)  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
778  | 
val OO_Grp =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
779  | 
let  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
780  | 
val map1 = Term.list_comb (mk_bnf_map Ds RTs As, map fst_const RTs);  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
781  | 
val map2 = Term.list_comb (mk_bnf_map Ds RTs Bs, map snd_const RTs);  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
782  | 
val bnf_in = mk_in setRs (map (mk_bnf_t Ds RTs) bnf_sets) CR;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
783  | 
in  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
784  | 
mk_rel_compp (mk_conversep (mk_Grp bnf_in map1), mk_Grp bnf_in map2)  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
785  | 
|> fold_rev Term.absfree Rs'  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
786  | 
end;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
787  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
788  | 
val rel_rhs = the_default OO_Grp rel_rhs_opt;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
789  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
790  | 
val rel_bind_def =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
791  | 
(fn () => def_qualify (if Binding.is_empty rel_b then mk_prefix_binding relN else rel_b),  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
792  | 
rel_rhs);  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
793  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
794  | 
val wit_rhss =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
795  | 
if null wit_rhss then  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
796  | 
[fold_rev Term.absdummy As (Term.list_comb (mk_bnf_map Ds As As,  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
797  | 
map2 (fn T => fn i => Term.absdummy T (Bound i)) As (live downto 1)) $  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
798  | 
          Const (@{const_name undefined}, CA))]
 | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
799  | 
else wit_rhss;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
800  | 
val nwits = length wit_rhss;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
801  | 
val wit_binds_defs =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
802  | 
let  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
803  | 
val bs = if nwits = 1 then [fn () => def_qualify (mk_prefix_binding witN)]  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
804  | 
else map (fn i => fn () => def_qualify (mk_prefix_binding (mk_witN i))) (1 upto nwits);  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
805  | 
in bs ~~ wit_rhss end;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
806  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
807  | 
val (((bnf_rel_term, raw_rel_def), (bnf_wit_terms, raw_wit_defs)), (lthy, lthy_old)) =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
808  | 
lthy  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
809  | 
|> maybe_define (is_some rel_rhs_opt) rel_bind_def  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
810  | 
||>> apfst split_list o fold_map (maybe_define (not (null wit_rhss))) wit_binds_defs  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
811  | 
||> `(maybe_restore lthy);  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
812  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
813  | 
val phi = Proof_Context.export_morphism lthy_old lthy;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
814  | 
val bnf_rel_def = Morphism.thm phi raw_rel_def;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
815  | 
val bnf_rel = Morphism.term phi bnf_rel_term;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
816  | 
fun mk_bnf_rel Ds As' Bs' =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
817  | 
normalize_rel lthy (map2 mk_pred2T As' Bs') (mk_bnf_T Ds As' Calpha) (mk_bnf_T Ds Bs' Calpha)  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
818  | 
bnf_rel;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
819  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
820  | 
val bnf_wit_defs = map (Morphism.thm phi) raw_wit_defs;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
821  | 
val bnf_wits =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
822  | 
map (normalize_wit Calpha_params Calpha alphas o Morphism.term phi) bnf_wit_terms;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
823  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
824  | 
fun mk_OO_Grp Ds' As' Bs' =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
825  | 
Term.subst_atomic_types ((Ds ~~ Ds') @ (As ~~ As') @ (Bs ~~ Bs')) OO_Grp;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
826  | 
in  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
827  | 
(((alphas, betas, deads, Calpha),  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
828  | 
(bnf_map, bnf_sets, bnf_bd, bnf_wits, bnf_rel),  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
829  | 
(bnf_map_def, bnf_set_defs, bnf_bd_def, bnf_wit_defs, bnf_rel_def),  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
830  | 
(mk_bnf_map, mk_bnf_t, mk_bnf_T, mk_bnf_rel, mk_OO_Grp)), lthy)  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
831  | 
end;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
832  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
833  | 
fun prepare_def const_policy mk_fact_policy qualify prep_typ prep_term Ds_opt map_b rel_b set_bs  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
834  | 
((((((raw_bnf_b, raw_bnf_T), raw_map), raw_sets), raw_bd), raw_wits), raw_rel_opt)  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
835  | 
no_defs_lthy =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
836  | 
let  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
837  | 
val fact_policy = mk_fact_policy no_defs_lthy;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
838  | 
val bnf_b = qualify raw_bnf_b;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
839  | 
val live = length raw_sets;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
840  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
841  | 
val T_rhs = prep_typ no_defs_lthy raw_bnf_T;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
842  | 
val map_rhs = prep_term no_defs_lthy raw_map;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
843  | 
val set_rhss = map (prep_term no_defs_lthy) raw_sets;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
844  | 
val bd_rhs = prep_term no_defs_lthy raw_bd;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
845  | 
val wit_rhss = map (prep_term no_defs_lthy) raw_wits;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
846  | 
val rel_rhs_opt = Option.map (prep_term no_defs_lthy) raw_rel_opt;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
847  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
848  | 
fun err T =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
849  | 
      error ("Trying to register the type " ^ quote (Syntax.string_of_typ no_defs_lthy T) ^
 | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
850  | 
" as unnamed BNF");  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
851  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
852  | 
val (bnf_b, key) =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
853  | 
if Binding.eq_name (bnf_b, Binding.empty) then  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
854  | 
(case T_rhs of  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
855  | 
Type (C, Ts) => if forall (can dest_TFree) Ts  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
856  | 
then (Binding.qualified_name C, C) else err T_rhs  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
857  | 
| T => err T)  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
858  | 
else (bnf_b, Local_Theory.full_name no_defs_lthy bnf_b);  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
859  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
860  | 
val (((alphas, betas, deads, Calpha),  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
861  | 
(bnf_map, bnf_sets, bnf_bd, bnf_wits, bnf_rel),  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
862  | 
(bnf_map_def, bnf_set_defs, bnf_bd_def, bnf_wit_defs, bnf_rel_def),  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
863  | 
(mk_bnf_map_Ds, mk_bnf_t_Ds, mk_bnf_T_Ds, _, mk_OO_Grp)), lthy) =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
864  | 
define_bnf_consts const_policy fact_policy Ds_opt map_b rel_b set_bs  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
865  | 
((((((bnf_b, T_rhs), map_rhs), set_rhss), bd_rhs), wit_rhss), rel_rhs_opt) no_defs_lthy;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
866  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
867  | 
val dead = length deads;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
868  | 
|
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
869  | 
val ((((((As', Bs'), Cs), Ds), B1Ts), B2Ts), (Ts, T)) = lthy  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
870  | 
|> mk_TFrees live  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
871  | 
||>> mk_TFrees live  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
872  | 
||>> mk_TFrees live  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
873  | 
||>> mk_TFrees dead  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
874  | 
||>> mk_TFrees live  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
875  | 
||>> mk_TFrees live  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
876  | 
||> fst o mk_TFrees 1  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
877  | 
||> the_single  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
878  | 
||> `(replicate live);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
879  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
880  | 
val mk_bnf_map = mk_bnf_map_Ds Ds;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
881  | 
val mk_bnf_t = mk_bnf_t_Ds Ds;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
882  | 
val mk_bnf_T = mk_bnf_T_Ds Ds;  | 
| 49453 | 883  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
884  | 
val pred2RTs = map2 mk_pred2T As' Bs';  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
885  | 
val pred2RTsAsCs = map2 mk_pred2T As' Cs;  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
886  | 
val pred2RTsBsCs = map2 mk_pred2T Bs' Cs;  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
887  | 
val pred2RT's = map2 mk_pred2T Bs' As';  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
888  | 
val self_pred2RTs = map2 mk_pred2T As' As';  | 
| 
52719
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
889  | 
val transfer_domRTs = map2 mk_pred2T As' B1Ts;  | 
| 
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
890  | 
val transfer_ranRTs = map2 mk_pred2T Bs' B2Ts;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
891  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
892  | 
val CA' = mk_bnf_T As' Calpha;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
893  | 
val CB' = mk_bnf_T Bs' Calpha;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
894  | 
val CC' = mk_bnf_T Cs Calpha;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
895  | 
val CB1 = mk_bnf_T B1Ts Calpha;  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
896  | 
val CB2 = mk_bnf_T B2Ts Calpha;  | 
| 49453 | 897  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
898  | 
val bnf_map_AsAs = mk_bnf_map As' As';  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
899  | 
val bnf_map_AsBs = mk_bnf_map As' Bs';  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
900  | 
val bnf_map_AsCs = mk_bnf_map As' Cs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
901  | 
val bnf_map_BsCs = mk_bnf_map Bs' Cs;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
902  | 
val bnf_sets_As = map (mk_bnf_t As') bnf_sets;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
903  | 
val bnf_sets_Bs = map (mk_bnf_t Bs') bnf_sets;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
904  | 
val bnf_bd_As = mk_bnf_t As' bnf_bd;  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
905  | 
fun mk_bnf_rel RTs CA CB = normalize_rel lthy RTs CA CB bnf_rel;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
906  | 
|
| 
49595
 
e8c57e59cbf8
got rid of other instance of shaky "Thm.generalize"
 
blanchet 
parents: 
49594 
diff
changeset
 | 
907  | 
val pre_names_lthy = lthy;  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
908  | 
val (((((((((((((((fs, gs), hs), x), y), zs), ys), As),  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
909  | 
As_copy), bs), Rs), Rs_copy), Ss),  | 
| 
52719
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
910  | 
transfer_domRs), transfer_ranRs), names_lthy) = pre_names_lthy  | 
| 52923 | 911  | 
|> mk_Frees "f" (map2 (curry op -->) As' Bs')  | 
912  | 
||>> mk_Frees "g" (map2 (curry op -->) Bs' Cs)  | 
|
913  | 
||>> mk_Frees "h" (map2 (curry op -->) As' Ts)  | 
|
| 51894 | 914  | 
||>> yield_singleton (mk_Frees "x") CA'  | 
915  | 
||>> yield_singleton (mk_Frees "y") CB'  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
916  | 
||>> mk_Frees "z" As'  | 
| 51916 | 917  | 
||>> mk_Frees "y" Bs'  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
918  | 
||>> mk_Frees "A" (map HOLogic.mk_setT As')  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
919  | 
||>> mk_Frees "A" (map HOLogic.mk_setT As')  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
920  | 
||>> mk_Frees "b" As'  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
921  | 
||>> mk_Frees "R" pred2RTs  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
922  | 
||>> mk_Frees "R" pred2RTs  | 
| 
52719
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
923  | 
||>> mk_Frees "S" pred2RTsBsCs  | 
| 
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
924  | 
||>> mk_Frees "R" transfer_domRTs  | 
| 
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
925  | 
||>> mk_Frees "S" transfer_ranRTs;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
926  | 
|
| 51762 | 927  | 
val fs_copy = map2 (retype_free o fastype_of) fs gs;  | 
928  | 
val x_copy = retype_free CA' y;  | 
|
929  | 
||
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
930  | 
val rel = mk_bnf_rel pred2RTs CA' CB';  | 
| 
52719
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
931  | 
val relAsAs = mk_bnf_rel self_pred2RTs CA' CA';  | 
| 
54189
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
932  | 
val bnf_wit_As = map (apsnd (mk_bnf_t As')) bnf_wits;  | 
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
933  | 
|
| 53270 | 934  | 
val map_id0_goal =  | 
| 51762 | 935  | 
let val bnf_map_app_id = Term.list_comb (bnf_map_AsAs, map HOLogic.id_const As') in  | 
936  | 
mk_Trueprop_eq (bnf_map_app_id, HOLogic.id_const CA')  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
937  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
938  | 
|
| 53287 | 939  | 
val map_comp0_goal =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
940  | 
let  | 
| 49018 | 941  | 
val bnf_map_app_comp = Term.list_comb (bnf_map_AsCs, map2 (curry HOLogic.mk_comp) gs fs);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
942  | 
val comp_bnf_map_app = HOLogic.mk_comp  | 
| 51762 | 943  | 
(Term.list_comb (bnf_map_BsCs, gs), Term.list_comb (bnf_map_AsBs, fs));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
944  | 
in  | 
| 
49123
 
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
 
blanchet 
parents: 
49111 
diff
changeset
 | 
945  | 
fold_rev Logic.all (fs @ gs) (mk_Trueprop_eq (bnf_map_app_comp, comp_bnf_map_app))  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
946  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
947  | 
|
| 51762 | 948  | 
fun mk_map_cong_prem x z set f f_copy =  | 
949  | 
Logic.all z (Logic.mk_implies  | 
|
950  | 
(HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set $ x)),  | 
|
951  | 
mk_Trueprop_eq (f $ z, f_copy $ z)));  | 
|
952  | 
||
| 
51761
 
4c9f08836d87
renamed "map_cong" axiom to "map_cong0" in preparation for real "map_cong"
 
blanchet 
parents: 
51758 
diff
changeset
 | 
953  | 
val map_cong0_goal =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
954  | 
let  | 
| 51762 | 955  | 
val prems = map4 (mk_map_cong_prem x) zs bnf_sets_As fs fs_copy;  | 
956  | 
val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x,  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
957  | 
Term.list_comb (bnf_map_AsBs, fs_copy) $ x);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
958  | 
in  | 
| 51762 | 959  | 
fold_rev Logic.all (x :: fs @ fs_copy) (Logic.list_implies (prems, eq))  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
960  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
961  | 
|
| 53289 | 962  | 
val set_map0s_goal =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
963  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
964  | 
fun mk_goal setA setB f =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
965  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
966  | 
val set_comp_map =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
967  | 
HOLogic.mk_comp (setB, Term.list_comb (bnf_map_AsBs, fs));  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
968  | 
val image_comp_set = HOLogic.mk_comp (mk_image f, setA);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
969  | 
in  | 
| 
49123
 
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
 
blanchet 
parents: 
49111 
diff
changeset
 | 
970  | 
fold_rev Logic.all fs (mk_Trueprop_eq (set_comp_map, image_comp_set))  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
971  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
972  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
973  | 
map3 mk_goal bnf_sets_As bnf_sets_Bs fs  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
974  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
975  | 
|
| 49458 | 976  | 
val card_order_bd_goal = HOLogic.mk_Trueprop (mk_card_order bnf_bd_As);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
977  | 
|
| 49458 | 978  | 
val cinfinite_bd_goal = HOLogic.mk_Trueprop (mk_cinfinite bnf_bd_As);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
979  | 
|
| 49458 | 980  | 
val set_bds_goal =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
981  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
982  | 
fun mk_goal set =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
983  | 
Logic.all x (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (set $ x)) bnf_bd_As));  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
984  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
985  | 
map mk_goal bnf_sets_As  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
986  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
987  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
988  | 
val relAsCs = mk_bnf_rel pred2RTsAsCs CA' CC';  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
989  | 
val relBsCs = mk_bnf_rel pred2RTsBsCs CB' CC';  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
990  | 
val rel_OO_lhs = Term.list_comb (relAsCs, map2 (curry mk_rel_compp) Rs Ss);  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
991  | 
val rel_OO_rhs = mk_rel_compp (Term.list_comb (rel, Rs), Term.list_comb (relBsCs, Ss));  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
992  | 
val le_rel_OO_goal =  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
993  | 
fold_rev Logic.all (Rs @ Ss) (HOLogic.mk_Trueprop (mk_leq rel_OO_rhs rel_OO_lhs));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
994  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
995  | 
val rel_OO_Grp_goal = fold_rev Logic.all Rs (mk_Trueprop_eq (Term.list_comb (rel, Rs),  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
996  | 
Term.list_comb (mk_OO_Grp Ds As' Bs', Rs)));  | 
| 49453 | 997  | 
|
| 53289 | 998  | 
val goals = zip_axioms map_id0_goal map_comp0_goal map_cong0_goal set_map0s_goal  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
999  | 
card_order_bd_goal cinfinite_bd_goal set_bds_goal le_rel_OO_goal rel_OO_Grp_goal;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1000  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1001  | 
fun mk_wit_goals (I, wit) =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1002  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1003  | 
val xs = map (nth bs) I;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1004  | 
fun wit_goal i =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1005  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1006  | 
val z = nth zs i;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1007  | 
val set_wit = nth bnf_sets_As i $ Term.list_comb (wit, xs);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1008  | 
val concl = HOLogic.mk_Trueprop  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1009  | 
(if member (op =) I i then HOLogic.mk_eq (z, nth bs i)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1010  | 
              else @{term False});
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1011  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1012  | 
fold_rev Logic.all (z :: xs)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1013  | 
(Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set_wit)), concl))  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1014  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1015  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1016  | 
map wit_goal (0 upto live - 1)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1017  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1018  | 
|
| 55197 | 1019  | 
fun triv_wit_tac ctxt = mk_trivial_wit_tac ctxt bnf_wit_defs;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1020  | 
|
| 
54189
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1021  | 
val wit_goalss =  | 
| 
54921
 
862c36b6e57c
avoid schematic variable in goal, which sometimes gets instantiated by tactic
 
blanchet 
parents: 
54841 
diff
changeset
 | 
1022  | 
(if null raw_wits then SOME triv_wit_tac else NONE, map mk_wit_goals bnf_wit_As);  | 
| 
54189
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1023  | 
|
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1024  | 
fun after_qed mk_wit_thms thms lthy =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1025  | 
let  | 
| 
54189
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1026  | 
val (axioms, nontriv_wit_thms) = apfst (mk_axioms live) (chop (length goals) thms);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1027  | 
|
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49103 
diff
changeset
 | 
1028  | 
        val bd_Card_order = #bd_card_order axioms RS @{thm conjunct2[OF card_order_on_Card_order]};
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1029  | 
        val bd_Cinfinite = @{thm conjI} OF [#bd_cinfinite axioms, bd_Card_order];
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1030  | 
        val bd_Cnotzero = bd_Cinfinite RS @{thm Cinfinite_Cnotzero};
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1031  | 
|
| 
51766
 
f19a4d0ab1bf
renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
 
blanchet 
parents: 
51765 
diff
changeset
 | 
1032  | 
fun mk_collect_set_map () =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1033  | 
let  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
1034  | 
val defT = mk_bnf_T Ts Calpha --> HOLogic.mk_setT T;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1035  | 
val collect_map = HOLogic.mk_comp  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1036  | 
(mk_collect (map (mk_bnf_t Ts) bnf_sets) defT,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1037  | 
Term.list_comb (mk_bnf_map As' Ts, hs));  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1038  | 
val image_collect = mk_collect  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1039  | 
(map2 (fn h => fn set => HOLogic.mk_comp (mk_image h, set)) hs bnf_sets_As)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1040  | 
defT;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1041  | 
            (*collect {set1 ... setm} o map f1 ... fm = collect {f1` o set1 ... fm` o setm}*)
 | 
| 
49123
 
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
 
blanchet 
parents: 
49111 
diff
changeset
 | 
1042  | 
val goal = fold_rev Logic.all hs (mk_Trueprop_eq (collect_map, image_collect));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1043  | 
in  | 
| 53289 | 1044  | 
Goal.prove_sorry lthy [] [] goal (K (mk_collect_set_map_tac (#set_map0 axioms)))  | 
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49103 
diff
changeset
 | 
1045  | 
|> Thm.close_derivation  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1046  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1047  | 
|
| 
51766
 
f19a4d0ab1bf
renamed "set_natural" to "set_map", reflecting {Bl,Po,Tr} concensus
 
blanchet 
parents: 
51765 
diff
changeset
 | 
1048  | 
val collect_set_map = Lazy.lazy mk_collect_set_map;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1049  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1050  | 
fun mk_in_mono () =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1051  | 
let  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1052  | 
val prems_mono = map2 (HOLogic.mk_Trueprop oo mk_leq) As As_copy;  | 
| 49458 | 1053  | 
val in_mono_goal =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1054  | 
fold_rev Logic.all (As @ As_copy)  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1055  | 
(Logic.list_implies (prems_mono, HOLogic.mk_Trueprop  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1056  | 
(mk_leq (mk_in As bnf_sets_As CA') (mk_in As_copy bnf_sets_As CA'))));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1057  | 
in  | 
| 51551 | 1058  | 
Goal.prove_sorry lthy [] [] in_mono_goal (K (mk_in_mono_tac live))  | 
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49103 
diff
changeset
 | 
1059  | 
|> Thm.close_derivation  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1060  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1061  | 
|
| 49538 | 1062  | 
val in_mono = Lazy.lazy mk_in_mono;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1063  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1064  | 
fun mk_in_cong () =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1065  | 
let  | 
| 51762 | 1066  | 
val prems_cong = map2 (curry mk_Trueprop_eq) As As_copy;  | 
| 49458 | 1067  | 
val in_cong_goal =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1068  | 
fold_rev Logic.all (As @ As_copy)  | 
| 51762 | 1069  | 
(Logic.list_implies (prems_cong,  | 
1070  | 
mk_Trueprop_eq (mk_in As bnf_sets_As CA', mk_in As_copy bnf_sets_As CA')));  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1071  | 
in  | 
| 51798 | 1072  | 
Goal.prove_sorry lthy [] [] in_cong_goal  | 
1073  | 
(K ((TRY o hyp_subst_tac lthy THEN' rtac refl) 1))  | 
|
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49103 
diff
changeset
 | 
1074  | 
|> Thm.close_derivation  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1075  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1076  | 
|
| 49538 | 1077  | 
val in_cong = Lazy.lazy mk_in_cong;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1078  | 
|
| 53285 | 1079  | 
val map_id = Lazy.lazy (fn () => mk_map_id (#map_id0 axioms));  | 
| 53288 | 1080  | 
val map_comp = Lazy.lazy (fn () => mk_map_comp (#map_comp0 axioms));  | 
| 51762 | 1081  | 
|
1082  | 
fun mk_map_cong () =  | 
|
1083  | 
let  | 
|
1084  | 
val prem0 = mk_Trueprop_eq (x, x_copy);  | 
|
1085  | 
val prems = map4 (mk_map_cong_prem x_copy) zs bnf_sets_As fs fs_copy;  | 
|
1086  | 
val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x,  | 
|
1087  | 
Term.list_comb (bnf_map_AsBs, fs_copy) $ x_copy);  | 
|
1088  | 
val goal = fold_rev Logic.all (x :: x_copy :: fs @ fs_copy)  | 
|
1089  | 
(Logic.list_implies (prem0 :: prems, eq));  | 
|
1090  | 
in  | 
|
| 51798 | 1091  | 
Goal.prove_sorry lthy [] [] goal (fn _ => mk_map_cong_tac lthy (#map_cong0 axioms))  | 
| 51762 | 1092  | 
|> Thm.close_derivation  | 
1093  | 
end;  | 
|
1094  | 
||
1095  | 
val map_cong = Lazy.lazy mk_map_cong;  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1096  | 
|
| 53290 | 1097  | 
val set_map = map (fn thm => Lazy.lazy (fn () => mk_set_map thm)) (#set_map0 axioms);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1098  | 
|
| 
54189
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1099  | 
val wit_thms =  | 
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1100  | 
if null nontriv_wit_thms then mk_wit_thms (map Lazy.force set_map) else nontriv_wit_thms;  | 
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1101  | 
|
| 
52635
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
1102  | 
fun mk_in_bd () =  | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
1103  | 
let  | 
| 52813 | 1104  | 
val bdT = fst (dest_relT (fastype_of bnf_bd_As));  | 
1105  | 
val bdTs = replicate live bdT;  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
1106  | 
val bd_bnfT = mk_bnf_T bdTs Calpha;  | 
| 52813 | 1107  | 
val surj_imp_ordLeq_inst = (if live = 0 then TrueI else  | 
1108  | 
let  | 
|
1109  | 
val ranTs = map (fn AT => mk_sumT (AT, HOLogic.unitT)) As';  | 
|
1110  | 
val funTs = map (fn T => bdT --> T) ranTs;  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
1111  | 
val ran_bnfT = mk_bnf_T ranTs Calpha;  | 
| 52813 | 1112  | 
val (revTs, Ts) = `rev (bd_bnfT :: funTs);  | 
1113  | 
val cTs = map (SOME o certifyT lthy) [ran_bnfT, Library.foldr1 HOLogic.mk_prodT Ts];  | 
|
1114  | 
val tinst = fold (fn T => fn t => HOLogic.mk_split (Term.absdummy T t)) (tl revTs)  | 
|
1115  | 
(Term.absdummy (hd revTs) (Term.list_comb (mk_bnf_map bdTs ranTs,  | 
|
1116  | 
map Bound (live - 1 downto 0)) $ Bound live));  | 
|
1117  | 
val cts = [NONE, SOME (certify lthy tinst)];  | 
|
1118  | 
in  | 
|
1119  | 
                Drule.instantiate' cTs cts @{thm surj_imp_ordLeq}
 | 
|
1120  | 
end);  | 
|
| 
52635
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
1121  | 
val bd = mk_cexp  | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
1122  | 
(if live = 0 then ctwo  | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
1123  | 
else mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo)  | 
| 52813 | 1124  | 
(mk_csum bnf_bd_As (mk_card_of (HOLogic.mk_UNIV bd_bnfT)));  | 
| 
52635
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
1125  | 
val in_bd_goal =  | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
1126  | 
fold_rev Logic.all As  | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
1127  | 
(HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (mk_in As bnf_sets_As CA')) bd));  | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
1128  | 
in  | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
1129  | 
Goal.prove_sorry lthy [] [] in_bd_goal  | 
| 55197 | 1130  | 
              (fn {context = ctxt, prems = _} => mk_in_bd_tac ctxt live surj_imp_ordLeq_inst
 | 
| 53288 | 1131  | 
(Lazy.force map_comp) (Lazy.force map_id) (#map_cong0 axioms)  | 
| 53290 | 1132  | 
(map Lazy.force set_map) (#set_bd axioms) (#bd_card_order axioms)  | 
| 
52635
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
1133  | 
bd_Card_order bd_Cinfinite bd_Cnotzero)  | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
1134  | 
|> Thm.close_derivation  | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
1135  | 
end;  | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
1136  | 
|
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
1137  | 
val in_bd = Lazy.lazy mk_in_bd;  | 
| 
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
1138  | 
|
| 53561 | 1139  | 
val rel_OO_Grp = #rel_OO_Grp axioms;  | 
1140  | 
val rel_OO_Grps = no_refl [rel_OO_Grp];  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1141  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1142  | 
fun mk_rel_Grp () =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1143  | 
let  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1144  | 
val lhs = Term.list_comb (rel, map2 mk_Grp As fs);  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1145  | 
val rhs = mk_Grp (mk_in As bnf_sets_As CA') (Term.list_comb (bnf_map_AsBs, fs));  | 
| 
49123
 
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
 
blanchet 
parents: 
49111 
diff
changeset
 | 
1146  | 
val goal = fold_rev Logic.all (As @ fs) (mk_Trueprop_eq (lhs, rhs));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1147  | 
in  | 
| 51551 | 1148  | 
Goal.prove_sorry lthy [] [] goal  | 
| 55197 | 1149  | 
              (fn {context = ctxt, prems = _} => mk_rel_Grp_tac ctxt rel_OO_Grps (#map_id0 axioms)
 | 
1150  | 
(#map_cong0 axioms) (Lazy.force map_id) (Lazy.force map_comp)  | 
|
1151  | 
(map Lazy.force set_map))  | 
|
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49103 
diff
changeset
 | 
1152  | 
|> Thm.close_derivation  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1153  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1154  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1155  | 
val rel_Grp = Lazy.lazy mk_rel_Grp;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1156  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1157  | 
fun mk_rel_prems f = map2 (HOLogic.mk_Trueprop oo f) Rs Rs_copy  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1158  | 
fun mk_rel_concl f = HOLogic.mk_Trueprop  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1159  | 
(f (Term.list_comb (rel, Rs), Term.list_comb (rel, Rs_copy)));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1160  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1161  | 
fun mk_rel_mono () =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1162  | 
let  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1163  | 
val mono_prems = mk_rel_prems mk_leq;  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1164  | 
val mono_concl = mk_rel_concl (uncurry mk_leq);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1165  | 
in  | 
| 51551 | 1166  | 
Goal.prove_sorry lthy [] []  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1167  | 
(fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (mono_prems, mono_concl)))  | 
| 
52844
 
66fa0f602cc4
more robust tactics (don't use unfolding when RHS might contain schematics not contained on the LHS)
 
traytel 
parents: 
52813 
diff
changeset
 | 
1168  | 
(K (mk_rel_mono_tac rel_OO_Grps (Lazy.force in_mono)))  | 
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49103 
diff
changeset
 | 
1169  | 
|> Thm.close_derivation  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1170  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1171  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1172  | 
fun mk_rel_cong () =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1173  | 
let  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1174  | 
val cong_prems = mk_rel_prems (curry HOLogic.mk_eq);  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1175  | 
val cong_concl = mk_rel_concl HOLogic.mk_eq;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1176  | 
in  | 
| 51551 | 1177  | 
Goal.prove_sorry lthy [] []  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1178  | 
(fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (cong_prems, cong_concl)))  | 
| 51798 | 1179  | 
(fn _ => (TRY o hyp_subst_tac lthy THEN' rtac refl) 1)  | 
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49103 
diff
changeset
 | 
1180  | 
|> Thm.close_derivation  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1181  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1182  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1183  | 
val rel_mono = Lazy.lazy mk_rel_mono;  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1184  | 
val rel_cong = Lazy.lazy mk_rel_cong;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1185  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1186  | 
fun mk_rel_eq () =  | 
| 
52719
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
1187  | 
Goal.prove_sorry lthy [] []  | 
| 
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
1188  | 
(mk_Trueprop_eq (Term.list_comb (relAsAs, map HOLogic.eq_const As'),  | 
| 
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
1189  | 
HOLogic.eq_const CA'))  | 
| 53270 | 1190  | 
(K (mk_rel_eq_tac live (Lazy.force rel_Grp) (Lazy.force rel_cong) (#map_id0 axioms)))  | 
| 
52719
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
1191  | 
|> Thm.close_derivation;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1192  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1193  | 
val rel_eq = Lazy.lazy mk_rel_eq;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1194  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1195  | 
fun mk_rel_conversep () =  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1196  | 
let  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1197  | 
val relBsAs = mk_bnf_rel pred2RT's CB' CA';  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1198  | 
val lhs = Term.list_comb (relBsAs, map mk_conversep Rs);  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1199  | 
val rhs = mk_conversep (Term.list_comb (rel, Rs));  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1200  | 
val le_goal = fold_rev Logic.all Rs (HOLogic.mk_Trueprop (mk_leq lhs rhs));  | 
| 51551 | 1201  | 
val le_thm = Goal.prove_sorry lthy [] [] le_goal  | 
| 55197 | 1202  | 
              (fn {context = ctxt, prems = _} => mk_rel_conversep_le_tac ctxt rel_OO_Grps
 | 
1203  | 
(Lazy.force rel_eq) (#map_cong0 axioms) (Lazy.force map_comp)  | 
|
1204  | 
(map Lazy.force set_map))  | 
|
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49103 
diff
changeset
 | 
1205  | 
|> Thm.close_derivation  | 
| 
49123
 
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
 
blanchet 
parents: 
49111 
diff
changeset
 | 
1206  | 
val goal = fold_rev Logic.all Rs (mk_Trueprop_eq (lhs, rhs));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1207  | 
in  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1208  | 
Goal.prove_sorry lthy [] [] goal  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1209  | 
(K (mk_rel_conversep_tac le_thm (Lazy.force rel_mono)))  | 
| 
49109
 
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
 
traytel 
parents: 
49103 
diff
changeset
 | 
1210  | 
|> Thm.close_derivation  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1211  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1212  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1213  | 
val rel_conversep = Lazy.lazy mk_rel_conversep;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1214  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1215  | 
fun mk_rel_OO () =  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
1216  | 
Goal.prove_sorry lthy [] []  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
1217  | 
(fold_rev Logic.all (Rs @ Ss) (HOLogic.mk_Trueprop (mk_leq rel_OO_lhs rel_OO_rhs)))  | 
| 55197 | 1218  | 
            (fn {context = ctxt, prems = _} => mk_rel_OO_le_tac ctxt rel_OO_Grps (Lazy.force rel_eq)
 | 
1219  | 
(#map_cong0 axioms) (Lazy.force map_comp) (map Lazy.force set_map))  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
1220  | 
|> Thm.close_derivation  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
1221  | 
          |> (fn thm => @{thm antisym} OF [thm, #le_rel_OO axioms]);
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1222  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1223  | 
val rel_OO = Lazy.lazy mk_rel_OO;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1224  | 
|
| 53561 | 1225  | 
        fun mk_in_rel () = trans OF [rel_OO_Grp, @{thm OO_Grp_alt}] RS @{thm predicate2_eqD};
 | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1226  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1227  | 
val in_rel = Lazy.lazy mk_in_rel;  | 
| 
49537
 
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
 
blanchet 
parents: 
49536 
diff
changeset
 | 
1228  | 
|
| 
 
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
 
blanchet 
parents: 
49536 
diff
changeset
 | 
1229  | 
fun mk_rel_flip () =  | 
| 
 
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
 
blanchet 
parents: 
49536 
diff
changeset
 | 
1230  | 
let  | 
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1231  | 
val rel_conversep_thm = Lazy.force rel_conversep;  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1232  | 
val cts = map (SOME o certify lthy) Rs;  | 
| 
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1233  | 
val rel_conversep_thm' = cterm_instantiate_pos cts rel_conversep_thm;  | 
| 
49537
 
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
 
blanchet 
parents: 
49536 
diff
changeset
 | 
1234  | 
in  | 
| 
51917
 
f964a9887713
store proper theorems even for fixed points that have no passive live variables
 
traytel 
parents: 
51916 
diff
changeset
 | 
1235  | 
            unfold_thms lthy @{thms conversep_iff} (rel_conversep_thm' RS @{thm predicate2_eqD})
 | 
| 
49595
 
e8c57e59cbf8
got rid of other instance of shaky "Thm.generalize"
 
blanchet 
parents: 
49594 
diff
changeset
 | 
1236  | 
|> singleton (Proof_Context.export names_lthy pre_names_lthy)  | 
| 
49537
 
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
 
blanchet 
parents: 
49536 
diff
changeset
 | 
1237  | 
end;  | 
| 
 
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
 
blanchet 
parents: 
49536 
diff
changeset
 | 
1238  | 
|
| 49538 | 1239  | 
val rel_flip = Lazy.lazy mk_rel_flip;  | 
| 
49537
 
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
 
blanchet 
parents: 
49536 
diff
changeset
 | 
1240  | 
|
| 51916 | 1241  | 
fun mk_rel_mono_strong () =  | 
1242  | 
let  | 
|
1243  | 
fun mk_prem setA setB R S a b =  | 
|
1244  | 
HOLogic.mk_Trueprop  | 
|
1245  | 
(mk_Ball (setA $ x) (Term.absfree (dest_Free a)  | 
|
1246  | 
(mk_Ball (setB $ y) (Term.absfree (dest_Free b)  | 
|
1247  | 
(HOLogic.mk_imp (R $ a $ b, S $ a $ b))))));  | 
|
1248  | 
val prems = HOLogic.mk_Trueprop (Term.list_comb (rel, Rs) $ x $ y) ::  | 
|
1249  | 
map6 mk_prem bnf_sets_As bnf_sets_Bs Rs Rs_copy zs ys;  | 
|
1250  | 
val concl = HOLogic.mk_Trueprop (Term.list_comb (rel, Rs_copy) $ x $ y);  | 
|
1251  | 
in  | 
|
1252  | 
Goal.prove_sorry lthy [] []  | 
|
1253  | 
(fold_rev Logic.all (x :: y :: Rs @ Rs_copy) (Logic.list_implies (prems, concl)))  | 
|
| 55197 | 1254  | 
              (fn {context = ctxt, prems = _} => mk_rel_mono_strong_tac ctxt (Lazy.force in_rel)
 | 
1255  | 
(map Lazy.force set_map))  | 
|
| 51916 | 1256  | 
|> Thm.close_derivation  | 
1257  | 
end;  | 
|
1258  | 
||
1259  | 
val rel_mono_strong = Lazy.lazy mk_rel_mono_strong;  | 
|
1260  | 
||
| 
52719
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
1261  | 
fun mk_map_transfer () =  | 
| 
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
1262  | 
let  | 
| 52725 | 1263  | 
val rels = map2 mk_fun_rel transfer_domRs transfer_ranRs;  | 
1264  | 
val rel = mk_fun_rel  | 
|
1265  | 
(Term.list_comb (mk_bnf_rel transfer_domRTs CA' CB1, transfer_domRs))  | 
|
1266  | 
(Term.list_comb (mk_bnf_rel transfer_ranRTs CB' CB2, transfer_ranRs));  | 
|
1267  | 
val concl = HOLogic.mk_Trueprop  | 
|
1268  | 
(fold_rev mk_fun_rel rels rel $ bnf_map_AsBs $ mk_bnf_map B1Ts B2Ts);  | 
|
| 
52719
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
1269  | 
in  | 
| 
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
1270  | 
Goal.prove_sorry lthy [] []  | 
| 52725 | 1271  | 
(fold_rev Logic.all (transfer_domRs @ transfer_ranRs) concl)  | 
| 55197 | 1272  | 
              (fn {context = ctxt, prems = _} => mk_map_transfer_tac ctxt (Lazy.force rel_mono)
 | 
1273  | 
(Lazy.force in_rel) (map Lazy.force set_map) (#map_cong0 axioms)  | 
|
1274  | 
(Lazy.force map_comp))  | 
|
| 
52719
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
1275  | 
|> Thm.close_derivation  | 
| 
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
1276  | 
end;  | 
| 
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
1277  | 
|
| 
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
1278  | 
val map_transfer = Lazy.lazy mk_map_transfer;  | 
| 
 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 
traytel 
parents: 
52635 
diff
changeset
 | 
1279  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1280  | 
val defs = mk_defs bnf_map_def bnf_set_defs bnf_rel_def;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1281  | 
|
| 
52635
 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 
traytel 
parents: 
51930 
diff
changeset
 | 
1282  | 
val facts = mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong  | 
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
1283  | 
in_mono in_rel map_comp map_cong map_id map_transfer rel_eq rel_flip set_map  | 
| 51916 | 1284  | 
rel_cong rel_mono rel_mono_strong rel_Grp rel_conversep rel_OO;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1285  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1286  | 
val wits = map2 mk_witness bnf_wits wit_thms;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1287  | 
|
| 49507 | 1288  | 
val bnf_rel =  | 
1289  | 
Term.subst_atomic_types ((Ds ~~ deads) @ (As' ~~ alphas) @ (Bs' ~~ betas)) rel;  | 
|
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1290  | 
|
| 
54841
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
1291  | 
val bnf = mk_bnf bnf_b Calpha live alphas betas dead deads bnf_map bnf_sets bnf_bd axioms  | 
| 
 
af71b753c459
express weak pullback property of bnfs only in terms of the relator
 
traytel 
parents: 
54740 
diff
changeset
 | 
1292  | 
defs facts wits bnf_rel;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1293  | 
in  | 
| 53265 | 1294  | 
(bnf, lthy |> note_bnf_thms fact_policy qualify bnf_b bnf)  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1295  | 
end;  | 
| 49459 | 1296  | 
|
1297  | 
val one_step_defs =  | 
|
| 
51893
 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 
traytel 
parents: 
51837 
diff
changeset
 | 
1298  | 
no_reflexive (bnf_map_def :: bnf_bd_def :: bnf_set_defs @ bnf_wit_defs @ [bnf_rel_def]);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1299  | 
in  | 
| 49459 | 1300  | 
(key, goals, wit_goalss, after_qed, lthy, one_step_defs)  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1301  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1302  | 
|
| 
49434
 
433dc7e028c8
separated registration of BNFs from bnf_def (BNFs are now stored only for bnf_def and (co)data commands)
 
traytel 
parents: 
49430 
diff
changeset
 | 
1303  | 
fun register_bnf key (bnf, lthy) =  | 
| 
54285
 
578371ba74cc
reverted 3e1d230f1c00 -- pervasiveness is useful, cf. Coinductive_Nat in the AFP
 
blanchet 
parents: 
54265 
diff
changeset
 | 
1304  | 
  (bnf, Local_Theory.declaration {syntax = false, pervasive = true}
 | 
| 
55444
 
ec73f81e49e7
iteration n in the 'default' vs. 'update_new' vs. 'update' saga -- 'update' makes sense now that we honor the canonical order on 'merge' (as opposed to raising 'DUP')
 
blanchet 
parents: 
55394 
diff
changeset
 | 
1305  | 
(fn phi => Data.map (Symtab.update (key, morph_bnf phi bnf))) lthy);  | 
| 
49434
 
433dc7e028c8
separated registration of BNFs from bnf_def (BNFs are now stored only for bnf_def and (co)data commands)
 
traytel 
parents: 
49430 
diff
changeset
 | 
1306  | 
|
| 
51767
 
bbcdd8519253
honor user-specified name for relator + generalize syntax
 
blanchet 
parents: 
51766 
diff
changeset
 | 
1307  | 
fun bnf_def const_policy fact_policy qualify tacs wit_tac Ds map_b rel_b set_bs =  | 
| 
54189
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1308  | 
(fn (_, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, one_step_defs) =>  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1309  | 
let  | 
| 55197 | 1310  | 
fun mk_wits_tac ctxt set_maps =  | 
1311  | 
TRYALL Goal.conjunction_tac THEN  | 
|
| 
54189
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1312  | 
(case triv_tac_opt of  | 
| 55197 | 1313  | 
SOME tac => tac ctxt set_maps  | 
1314  | 
| NONE => unfold_thms_tac ctxt one_step_defs THEN wit_tac ctxt);  | 
|
| 
54189
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1315  | 
val wit_goals = map Logic.mk_conjunction_balanced wit_goalss;  | 
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1316  | 
fun mk_wit_thms set_maps =  | 
| 55197 | 1317  | 
Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals)  | 
1318  | 
        (fn {context = ctxt, prems = _} => mk_wits_tac ctxt set_maps)
 | 
|
| 
54189
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1319  | 
|> Conjunction.elim_balanced (length wit_goals)  | 
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1320  | 
|> map2 (Conjunction.elim_balanced o length) wit_goalss  | 
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1321  | 
|> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1322  | 
in  | 
| 51551 | 1323  | 
map2 (Thm.close_derivation oo Goal.prove_sorry lthy [] [])  | 
| 55197 | 1324  | 
      goals (map (fn tac => fn {context = ctxt, prems = _} =>
 | 
1325  | 
unfold_thms_tac ctxt one_step_defs THEN tac ctxt) tacs)  | 
|
| 
54189
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1326  | 
|> (fn thms => after_qed mk_wit_thms (map single thms) lthy)  | 
| 54421 | 1327  | 
end) oo prepare_def const_policy fact_policy qualify (K I) (K I) Ds map_b rel_b set_bs;  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1328  | 
|
| 
54189
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1329  | 
val bnf_cmd = (fn (key, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, defs) =>  | 
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1330  | 
let  | 
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1331  | 
val wit_goals = map Logic.mk_conjunction_balanced wit_goalss;  | 
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1332  | 
fun mk_triv_wit_thms tac set_maps =  | 
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1333  | 
Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals)  | 
| 55197 | 1334  | 
        (fn {context = ctxt, prems = _} => TRYALL Goal.conjunction_tac THEN tac ctxt set_maps)
 | 
| 
54189
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1335  | 
|> Conjunction.elim_balanced (length wit_goals)  | 
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1336  | 
|> map2 (Conjunction.elim_balanced o length) wit_goalss  | 
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1337  | 
|> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));  | 
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1338  | 
val (mk_wit_thms, nontriv_wit_goals) =  | 
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1339  | 
(case triv_tac_opt of  | 
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1340  | 
NONE => (fn _ => [], map (map (rpair [])) wit_goalss)  | 
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1341  | 
| SOME tac => (mk_triv_wit_thms tac, []));  | 
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1342  | 
in  | 
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1343  | 
Proof.unfolding ([[(defs, [])]])  | 
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1344  | 
(Proof.theorem NONE (snd o register_bnf key oo after_qed mk_wit_thms)  | 
| 
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1345  | 
(map (single o rpair []) goals @ nontriv_wit_goals) lthy)  | 
| 54421 | 1346  | 
end) oo prepare_def Do_Inline (user_policy Note_Some) I Syntax.read_typ Syntax.read_term NONE  | 
| 
54189
 
c0186a0d8cb3
define a trivial nonemptiness witness if none is provided
 
traytel 
parents: 
54158 
diff
changeset
 | 
1347  | 
Binding.empty Binding.empty [];  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1348  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1349  | 
fun print_bnfs ctxt =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1350  | 
let  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1351  | 
fun pretty_set sets i = Pretty.block  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1352  | 
[Pretty.str (mk_setN (i + 1) ^ ":"), Pretty.brk 1,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1353  | 
Pretty.quote (Syntax.pretty_term ctxt (nth sets i))];  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1354  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1355  | 
    fun pretty_bnf (key, BNF {T = T, map = map, sets = sets, bd = bd,
 | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1356  | 
live = live, lives = lives, dead = dead, deads = deads, ...}) =  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1357  | 
Pretty.big_list  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1358  | 
(Pretty.string_of (Pretty.block [Pretty.str key, Pretty.str ":", Pretty.brk 1,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1359  | 
Pretty.quote (Syntax.pretty_typ ctxt T)]))  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1360  | 
([Pretty.block [Pretty.str "live:", Pretty.brk 1, Pretty.str (string_of_int live),  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1361  | 
Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) lives)],  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1362  | 
Pretty.block [Pretty.str "dead:", Pretty.brk 1, Pretty.str (string_of_int dead),  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1363  | 
Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) deads)],  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1364  | 
Pretty.block [Pretty.str (mapN ^ ":"), Pretty.brk 1,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1365  | 
Pretty.quote (Syntax.pretty_term ctxt map)]] @  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1366  | 
List.map (pretty_set sets) (0 upto length sets - 1) @  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1367  | 
[Pretty.block [Pretty.str (bdN ^ ":"), Pretty.brk 1,  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1368  | 
Pretty.quote (Syntax.pretty_term ctxt bd)]]);  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1369  | 
in  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1370  | 
Pretty.big_list "BNFs:" (map pretty_bnf (Symtab.dest (Data.get (Context.Proof ctxt))))  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1371  | 
|> Pretty.writeln  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1372  | 
end;  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1373  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1374  | 
val _ =  | 
| 
51836
 
4d6dcd51dd52
renamed "bnf_def" keyword to "bnf" (since it's not a definition, but rather a registration)
 
blanchet 
parents: 
51823 
diff
changeset
 | 
1375  | 
  Outer_Syntax.improper_command @{command_spec "print_bnfs"}
 | 
| 53289 | 1376  | 
"print all bounded natural functors"  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1377  | 
(Scan.succeed (Toplevel.keep (print_bnfs o Toplevel.context_of)));  | 
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1378  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1379  | 
val _ =  | 
| 
51836
 
4d6dcd51dd52
renamed "bnf_def" keyword to "bnf" (since it's not a definition, but rather a registration)
 
blanchet 
parents: 
51823 
diff
changeset
 | 
1380  | 
  Outer_Syntax.local_theory_to_proof @{command_spec "bnf"}
 | 
| 53289 | 1381  | 
"register a type as a bounded natural functor"  | 
| 54421 | 1382  | 
(parse_opt_binding_colon -- Parse.typ --|  | 
1383  | 
       (Parse.reserved "map" -- @{keyword ":"}) -- Parse.term --
 | 
|
1384  | 
       (Scan.option ((Parse.reserved "sets" -- @{keyword ":"}) |--
 | 
|
1385  | 
Scan.repeat1 (Scan.unless (Parse.reserved "bd") Parse.term)) >> the_default []) --|  | 
|
1386  | 
       (Parse.reserved "bd" -- @{keyword ":"}) -- Parse.term --
 | 
|
1387  | 
       (Scan.option ((Parse.reserved "wits" -- @{keyword ":"}) |--
 | 
|
1388  | 
Scan.repeat1 (Scan.unless (Parse.reserved "rel") Parse.term)) >> the_default []) --  | 
|
1389  | 
       Scan.option ((Parse.reserved "rel" -- @{keyword ":"}) |-- Parse.term)
 | 
|
| 
51836
 
4d6dcd51dd52
renamed "bnf_def" keyword to "bnf" (since it's not a definition, but rather a registration)
 
blanchet 
parents: 
51823 
diff
changeset
 | 
1390  | 
>> bnf_cmd);  | 
| 
48975
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1391  | 
|
| 
 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 
blanchet 
parents:  
diff
changeset
 | 
1392  | 
end;  |