| author | blanchet |
| Fri, 08 Mar 2013 14:15:39 +0100 | |
| changeset 51380 | cac8c9a636b6 |
| parent 49631 | 9ce0c2cbadb8 |
| child 51551 | 88d1d19fb74f |
| permissions | -rw-r--r-- |
|
49509
163914705f8d
renamed top-level theory from "Codatatype" to "BNF"
blanchet
parents:
49507
diff
changeset
|
1 |
(* Title: HOL/BNF/Tools/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 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
11 |
type BNF |
|
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 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
14 |
val bnf_of: Proof.context -> string -> BNF option |
|
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
|
15 |
val register_bnf: string -> (BNF * local_theory) -> (BNF * local_theory) |
|
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
|
16 |
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
17 |
val name_of_bnf: BNF -> binding |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
18 |
val T_of_bnf: BNF -> typ |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
19 |
val live_of_bnf: BNF -> int |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
20 |
val lives_of_bnf: BNF -> typ list |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
21 |
val dead_of_bnf: BNF -> int |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
22 |
val deads_of_bnf: BNF -> typ list |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
23 |
val nwits_of_bnf: BNF -> int |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
24 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
25 |
val mapN: string |
| 49507 | 26 |
val relN: string |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
27 |
val setN: string |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
28 |
val mk_setN: int -> string |
| 49506 | 29 |
val srelN: string |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
30 |
|
|
49214
2a3cb4c71b87
construct the right iterator theorem in the recursive case
blanchet
parents:
49123
diff
changeset
|
31 |
val map_of_bnf: BNF -> term |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49538
diff
changeset
|
32 |
val sets_of_bnf: BNF -> term list |
|
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49538
diff
changeset
|
33 |
val rel_of_bnf: BNF -> term |
|
49214
2a3cb4c71b87
construct the right iterator theorem in the recursive case
blanchet
parents:
49123
diff
changeset
|
34 |
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
35 |
val mk_T_of_bnf: typ list -> typ list -> BNF -> typ |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
36 |
val mk_bd_of_bnf: typ list -> typ list -> BNF -> term |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
37 |
val mk_map_of_bnf: typ list -> typ list -> typ list -> BNF -> term |
| 49507 | 38 |
val mk_rel_of_bnf: typ list -> typ list -> typ list -> BNF -> term |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
39 |
val mk_sets_of_bnf: typ list list -> typ list list -> BNF -> term list |
| 49506 | 40 |
val mk_srel_of_bnf: typ list -> typ list -> typ list -> BNF -> term |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
41 |
val mk_wits_of_bnf: typ list list -> typ list list -> BNF -> (int list * term) list |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
42 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
43 |
val bd_Card_order_of_bnf: BNF -> thm |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
44 |
val bd_Cinfinite_of_bnf: BNF -> thm |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
45 |
val bd_Cnotzero_of_bnf: BNF -> thm |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
46 |
val bd_card_order_of_bnf: BNF -> thm |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
47 |
val bd_cinfinite_of_bnf: BNF -> thm |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
48 |
val collect_set_natural_of_bnf: BNF -> thm |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
49 |
val in_bd_of_bnf: BNF -> thm |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
50 |
val in_cong_of_bnf: BNF -> thm |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
51 |
val in_mono_of_bnf: BNF -> thm |
| 49506 | 52 |
val in_srel_of_bnf: BNF -> thm |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
53 |
val map_comp'_of_bnf: BNF -> thm |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
54 |
val map_comp_of_bnf: BNF -> thm |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
55 |
val map_cong_of_bnf: BNF -> thm |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
56 |
val map_def_of_bnf: BNF -> thm |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
57 |
val map_id'_of_bnf: BNF -> thm |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
58 |
val map_id_of_bnf: BNF -> thm |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
59 |
val map_wppull_of_bnf: BNF -> thm |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
60 |
val map_wpull_of_bnf: BNF -> thm |
| 49507 | 61 |
val rel_def_of_bnf: BNF -> thm |
|
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49585
diff
changeset
|
62 |
val rel_eq_of_bnf: BNF -> thm |
|
49537
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
63 |
val rel_flip_of_bnf: BNF -> thm |
|
49594
55e798614c45
tweaked theorem names (in particular, dropped s's)
blanchet
parents:
49592
diff
changeset
|
64 |
val rel_srel_of_bnf: BNF -> thm |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
65 |
val set_bd_of_bnf: BNF -> thm list |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
66 |
val set_defs_of_bnf: BNF -> thm list |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
67 |
val set_natural'_of_bnf: BNF -> thm list |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
68 |
val set_natural_of_bnf: BNF -> thm list |
| 49506 | 69 |
val srel_def_of_bnf: BNF -> thm |
70 |
val srel_Gr_of_bnf: BNF -> thm |
|
71 |
val srel_Id_of_bnf: BNF -> thm |
|
72 |
val srel_O_of_bnf: BNF -> thm |
|
73 |
val srel_O_Gr_of_bnf: BNF -> thm |
|
74 |
val srel_cong_of_bnf: BNF -> thm |
|
75 |
val srel_converse_of_bnf: BNF -> thm |
|
76 |
val srel_mono_of_bnf: BNF -> thm |
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
77 |
val wit_thms_of_bnf: BNF -> thm list |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
78 |
val wit_thmss_of_bnf: BNF -> thm list list |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
79 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
80 |
val mk_witness: int list * term -> thm list -> nonemptiness_witness |
| 49103 | 81 |
val minimize_wits: (''a list * 'b) list -> (''a list * 'b) list
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
82 |
val wits_of_bnf: BNF -> nonemptiness_witness list |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
83 |
|
|
49460
4dd451a075ce
fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents:
49459
diff
changeset
|
84 |
val zip_axioms: 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a -> 'a list |
| 49456 | 85 |
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
86 |
datatype const_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline |
| 49538 | 87 |
datatype fact_policy = Dont_Note | Note_Some | Note_All |
88 |
||
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
89 |
val bnf_note_all: bool Config.T |
| 49435 | 90 |
val user_policy: fact_policy -> Proof.context -> fact_policy |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
91 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
92 |
val print_bnfs: Proof.context -> unit |
| 49018 | 93 |
val bnf_def: const_policy -> (Proof.context -> fact_policy) -> (binding -> binding) -> |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
94 |
({prems: thm list, context: Proof.context} -> tactic) list ->
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
95 |
({prems: thm list, context: Proof.context} -> tactic) -> typ list option ->
|
| 49459 | 96 |
((((binding * term) * term list) * term) * term list) * term option -> local_theory -> |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
97 |
BNF * local_theory |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
98 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
99 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
100 |
structure BNF_Def : BNF_DEF = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
101 |
struct |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
102 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
103 |
open BNF_Util |
| 49463 | 104 |
open BNF_Tactics |
| 49284 | 105 |
open BNF_Def_Tactics |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
106 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
107 |
type axioms = {
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
108 |
map_id: thm, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
109 |
map_comp: thm, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
110 |
map_cong: thm, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
111 |
set_natural: thm list, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
112 |
bd_card_order: thm, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
113 |
bd_cinfinite: thm, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
114 |
set_bd: thm list, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
115 |
in_bd: thm, |
| 49453 | 116 |
map_wpull: thm, |
| 49506 | 117 |
srel_O_Gr: thm |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
118 |
}; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
119 |
|
| 49506 | 120 |
fun mk_axioms' (((((((((id, comp), cong), nat), c_o), cinf), set_bd), in_bd), wpull), srel) = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
121 |
{map_id = id, map_comp = comp, map_cong = cong, set_natural = nat, bd_card_order = c_o,
|
| 49506 | 122 |
bd_cinfinite = cinf, set_bd = set_bd, in_bd = in_bd, map_wpull = wpull, srel_O_Gr = srel}; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
123 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
124 |
fun dest_cons [] = raise Empty |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
125 |
| 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
|
126 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
127 |
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
|
128 |
|> map the_single |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
129 |
|> dest_cons |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
130 |
||>> dest_cons |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
131 |
||>> dest_cons |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
132 |
||>> chop n |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
133 |
||>> dest_cons |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
134 |
||>> dest_cons |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
135 |
||>> chop n |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
136 |
||>> dest_cons |
| 49453 | 137 |
||>> dest_cons |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
138 |
||> the_single |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
139 |
|> mk_axioms'; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
140 |
|
| 49506 | 141 |
fun zip_axioms mid mcomp mcong snat bdco bdinf sbd inbd wpull srel = |
142 |
[mid, mcomp, mcong] @ snat @ [bdco, bdinf] @ sbd @ [inbd, wpull, srel]; |
|
|
49460
4dd451a075ce
fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents:
49459
diff
changeset
|
143 |
|
| 49453 | 144 |
fun dest_axioms {map_id, map_comp, map_cong, set_natural, bd_card_order, bd_cinfinite, set_bd,
|
| 49506 | 145 |
in_bd, map_wpull, srel_O_Gr} = |
|
49460
4dd451a075ce
fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents:
49459
diff
changeset
|
146 |
zip_axioms map_id map_comp map_cong set_natural bd_card_order bd_cinfinite set_bd in_bd map_wpull |
| 49506 | 147 |
srel_O_Gr; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
148 |
|
| 49463 | 149 |
fun map_axioms f {map_id, map_comp, map_cong, set_natural, bd_card_order, bd_cinfinite, set_bd,
|
| 49506 | 150 |
in_bd, map_wpull, srel_O_Gr} = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
151 |
{map_id = f map_id,
|
| 49463 | 152 |
map_comp = f map_comp, |
153 |
map_cong = f map_cong, |
|
154 |
set_natural = map f set_natural, |
|
155 |
bd_card_order = f bd_card_order, |
|
156 |
bd_cinfinite = f bd_cinfinite, |
|
157 |
set_bd = map f set_bd, |
|
158 |
in_bd = f in_bd, |
|
159 |
map_wpull = f map_wpull, |
|
| 49506 | 160 |
srel_O_Gr = f srel_O_Gr}; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
161 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
162 |
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
|
163 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
164 |
type defs = {
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
165 |
map_def: thm, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
166 |
set_defs: thm list, |
| 49507 | 167 |
rel_def: thm, |
| 49506 | 168 |
srel_def: thm |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
169 |
} |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
170 |
|
| 49507 | 171 |
fun mk_defs map sets rel srel = {map_def = map, set_defs = sets, rel_def = rel, srel_def = srel};
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
172 |
|
| 49507 | 173 |
fun map_defs f {map_def, set_defs, rel_def, srel_def} =
|
174 |
{map_def = f map_def, set_defs = map f set_defs, rel_def = f rel_def, srel_def = f srel_def};
|
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
175 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
176 |
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
|
177 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
178 |
type facts = {
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
179 |
bd_Card_order: thm, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
180 |
bd_Cinfinite: thm, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
181 |
bd_Cnotzero: thm, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
182 |
collect_set_natural: thm lazy, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
183 |
in_cong: thm lazy, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
184 |
in_mono: thm lazy, |
| 49506 | 185 |
in_srel: thm lazy, |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
186 |
map_comp': thm lazy, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
187 |
map_id': thm lazy, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
188 |
map_wppull: thm lazy, |
|
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49585
diff
changeset
|
189 |
rel_eq: thm lazy, |
|
49537
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
190 |
rel_flip: thm lazy, |
|
49594
55e798614c45
tweaked theorem names (in particular, dropped s's)
blanchet
parents:
49592
diff
changeset
|
191 |
rel_srel: thm lazy, |
| 49506 | 192 |
set_natural': thm lazy list, |
193 |
srel_cong: thm lazy, |
|
194 |
srel_mono: thm lazy, |
|
195 |
srel_Id: thm lazy, |
|
196 |
srel_Gr: thm lazy, |
|
197 |
srel_converse: thm lazy, |
|
198 |
srel_O: thm lazy |
|
|
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 |
|
| 49506 | 201 |
fun mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_natural in_cong in_mono in_srel |
|
49594
55e798614c45
tweaked theorem names (in particular, dropped s's)
blanchet
parents:
49592
diff
changeset
|
202 |
map_comp' map_id' map_wppull rel_eq rel_flip rel_srel set_natural' srel_cong srel_mono |
|
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49585
diff
changeset
|
203 |
srel_Id srel_Gr srel_converse srel_O = {
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
204 |
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
|
205 |
bd_Cinfinite = bd_Cinfinite, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
206 |
bd_Cnotzero = bd_Cnotzero, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
207 |
collect_set_natural = collect_set_natural, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
208 |
in_cong = in_cong, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
209 |
in_mono = in_mono, |
| 49506 | 210 |
in_srel = in_srel, |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
211 |
map_comp' = map_comp', |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
212 |
map_id' = map_id', |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
213 |
map_wppull = map_wppull, |
|
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49585
diff
changeset
|
214 |
rel_eq = rel_eq, |
|
49537
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
215 |
rel_flip = rel_flip, |
|
49594
55e798614c45
tweaked theorem names (in particular, dropped s's)
blanchet
parents:
49592
diff
changeset
|
216 |
rel_srel = rel_srel, |
| 49506 | 217 |
set_natural' = set_natural', |
218 |
srel_cong = srel_cong, |
|
219 |
srel_mono = srel_mono, |
|
220 |
srel_Id = srel_Id, |
|
221 |
srel_Gr = srel_Gr, |
|
222 |
srel_converse = srel_converse, |
|
223 |
srel_O = srel_O}; |
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
224 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
225 |
fun map_facts f {
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
226 |
bd_Card_order, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
227 |
bd_Cinfinite, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
228 |
bd_Cnotzero, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
229 |
collect_set_natural, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
230 |
in_cong, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
231 |
in_mono, |
| 49506 | 232 |
in_srel, |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
233 |
map_comp', |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
234 |
map_id', |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
235 |
map_wppull, |
|
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49585
diff
changeset
|
236 |
rel_eq, |
|
49537
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
237 |
rel_flip, |
|
49594
55e798614c45
tweaked theorem names (in particular, dropped s's)
blanchet
parents:
49592
diff
changeset
|
238 |
rel_srel, |
| 49506 | 239 |
set_natural', |
240 |
srel_cong, |
|
241 |
srel_mono, |
|
242 |
srel_Id, |
|
243 |
srel_Gr, |
|
244 |
srel_converse, |
|
245 |
srel_O} = |
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
246 |
{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
|
247 |
bd_Cinfinite = f bd_Cinfinite, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
248 |
bd_Cnotzero = f bd_Cnotzero, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
249 |
collect_set_natural = Lazy.map f collect_set_natural, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
250 |
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
|
251 |
in_mono = Lazy.map f in_mono, |
| 49506 | 252 |
in_srel = Lazy.map f in_srel, |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
253 |
map_comp' = Lazy.map f map_comp', |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
254 |
map_id' = Lazy.map f map_id', |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
255 |
map_wppull = Lazy.map f map_wppull, |
|
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49585
diff
changeset
|
256 |
rel_eq = Lazy.map f rel_eq, |
|
49537
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
257 |
rel_flip = Lazy.map f rel_flip, |
|
49594
55e798614c45
tweaked theorem names (in particular, dropped s's)
blanchet
parents:
49592
diff
changeset
|
258 |
rel_srel = Lazy.map f rel_srel, |
| 49506 | 259 |
set_natural' = map (Lazy.map f) set_natural', |
260 |
srel_cong = Lazy.map f srel_cong, |
|
261 |
srel_mono = Lazy.map f srel_mono, |
|
262 |
srel_Id = Lazy.map f srel_Id, |
|
263 |
srel_Gr = Lazy.map f srel_Gr, |
|
264 |
srel_converse = Lazy.map f srel_converse, |
|
265 |
srel_O = Lazy.map f srel_O}; |
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
266 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
267 |
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
|
268 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
269 |
type nonemptiness_witness = {
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
270 |
I: int list, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
271 |
wit: term, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
272 |
prop: thm list |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
273 |
}; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
274 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
275 |
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
|
276 |
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
|
277 |
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
|
278 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
279 |
datatype BNF = BNF of {
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
280 |
name: binding, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
281 |
T: typ, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
282 |
live: int, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
283 |
lives: typ list, (*source type variables of map, only for composition*) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
284 |
lives': typ list, (*target type variables of map, only for composition*) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
285 |
dead: int, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
286 |
deads: typ list, (*only for composition*) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
287 |
map: term, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
288 |
sets: term list, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
289 |
bd: term, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
290 |
axioms: axioms, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
291 |
defs: defs, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
292 |
facts: facts, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
293 |
nwits: int, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
294 |
wits: nonemptiness_witness list, |
| 49507 | 295 |
rel: term, |
| 49506 | 296 |
srel: term |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
297 |
}; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
298 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
299 |
(* getters *) |
|
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 |
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
|
302 |
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
|
303 |
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
|
304 |
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
|
305 |
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
|
306 |
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
|
307 |
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
|
308 |
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
|
309 |
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
|
310 |
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
|
311 |
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
|
312 |
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
|
313 |
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
|
314 |
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
|
315 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
316 |
(*terms*) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
317 |
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
|
318 |
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
|
319 |
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
|
320 |
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
|
321 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
322 |
Term.subst_atomic_types |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
323 |
((#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
|
324 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
325 |
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
|
326 |
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
|
327 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
328 |
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
|
329 |
((#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
|
330 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
331 |
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
|
332 |
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
|
333 |
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
|
334 |
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
|
335 |
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
|
336 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
337 |
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
|
338 |
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
|
339 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
340 |
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
|
341 |
((#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
|
342 |
end; |
| 49507 | 343 |
val rel_of_bnf = #rel o rep_bnf; |
344 |
fun mk_rel_of_bnf Ds Ts Us bnf = |
|
| 49462 | 345 |
let val bnf_rep = rep_bnf bnf; |
346 |
in |
|
347 |
Term.subst_atomic_types |
|
| 49507 | 348 |
((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#rel bnf_rep) |
| 49462 | 349 |
end; |
| 49506 | 350 |
val srel_of_bnf = #srel o rep_bnf; |
351 |
fun mk_srel_of_bnf Ds Ts Us bnf = |
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
352 |
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
|
353 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
354 |
Term.subst_atomic_types |
| 49506 | 355 |
((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#srel bnf_rep) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
356 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
357 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
358 |
(*thms*) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
359 |
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
|
360 |
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
|
361 |
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
|
362 |
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
|
363 |
val bd_Cnotzero_of_bnf = #bd_Cnotzero o #facts o rep_bnf; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
364 |
val collect_set_natural_of_bnf = Lazy.force o #collect_set_natural o #facts o rep_bnf; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
365 |
val in_bd_of_bnf = #in_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
|
366 |
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
|
367 |
val in_mono_of_bnf = Lazy.force o #in_mono o #facts o rep_bnf; |
| 49506 | 368 |
val in_srel_of_bnf = Lazy.force o #in_srel 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
|
369 |
val map_def_of_bnf = #map_def o #defs o rep_bnf; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
370 |
val map_id_of_bnf = #map_id o #axioms o rep_bnf; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
371 |
val map_id'_of_bnf = Lazy.force o #map_id' o #facts o rep_bnf; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
372 |
val map_comp_of_bnf = #map_comp o #axioms o rep_bnf; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
373 |
val map_comp'_of_bnf = Lazy.force o #map_comp' o #facts o rep_bnf; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
374 |
val map_cong_of_bnf = #map_cong o #axioms o rep_bnf; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
375 |
val map_wppull_of_bnf = Lazy.force o #map_wppull o #facts o rep_bnf; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
376 |
val map_wpull_of_bnf = #map_wpull o #axioms o rep_bnf; |
| 49507 | 377 |
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
|
378 |
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
|
379 |
val rel_flip_of_bnf = Lazy.force o #rel_flip o #facts o rep_bnf; |
|
49594
55e798614c45
tweaked theorem names (in particular, dropped s's)
blanchet
parents:
49592
diff
changeset
|
380 |
val rel_srel_of_bnf = Lazy.force o #rel_srel 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
|
381 |
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
|
382 |
val set_defs_of_bnf = #set_defs o #defs o rep_bnf; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
383 |
val set_natural_of_bnf = #set_natural o #axioms o rep_bnf; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
384 |
val set_natural'_of_bnf = map Lazy.force o #set_natural' o #facts o rep_bnf; |
| 49506 | 385 |
val srel_cong_of_bnf = Lazy.force o #srel_cong o #facts o rep_bnf; |
386 |
val srel_mono_of_bnf = Lazy.force o #srel_mono o #facts o rep_bnf; |
|
387 |
val srel_def_of_bnf = #srel_def o #defs o rep_bnf; |
|
388 |
val srel_Id_of_bnf = Lazy.force o #srel_Id o #facts o rep_bnf; |
|
389 |
val srel_Gr_of_bnf = Lazy.force o #srel_Gr o #facts o rep_bnf; |
|
390 |
val srel_converse_of_bnf = Lazy.force o #srel_converse o #facts o rep_bnf; |
|
391 |
val srel_O_of_bnf = Lazy.force o #srel_O o #facts o rep_bnf; |
|
392 |
val srel_O_Gr_of_bnf = #srel_O_Gr 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
|
393 |
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
|
394 |
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
|
395 |
|
| 49507 | 396 |
fun mk_bnf name T live lives lives' dead deads map sets bd axioms defs facts wits rel srel = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
397 |
BNF {name = name, T = T,
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
398 |
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
|
399 |
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
|
400 |
axioms = axioms, defs = defs, facts = facts, |
| 49507 | 401 |
nwits = length wits, wits = wits, rel = rel, srel = srel}; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
402 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
403 |
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
|
404 |
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
|
405 |
axioms = axioms, defs = defs, facts = facts, |
| 49507 | 406 |
nwits = nwits, wits = wits, rel = rel, srel = srel}) = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
407 |
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
|
408 |
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
|
409 |
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
|
410 |
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
|
411 |
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
|
412 |
bd = Morphism.term phi bd, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
413 |
axioms = morph_axioms phi axioms, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
414 |
defs = morph_defs phi defs, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
415 |
facts = morph_facts phi facts, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
416 |
nwits = nwits, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
417 |
wits = List.map (morph_witness phi) wits, |
| 49507 | 418 |
rel = Morphism.term phi rel, srel = Morphism.term phi srel}; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
419 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
420 |
fun eq_bnf (BNF {T = T1, live = live1, dead = dead1, ...},
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
421 |
BNF {T = T2, live = live2, dead = dead2, ...}) =
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
422 |
Type.could_unify (T1, T2) andalso live1 = live2 andalso dead1 = dead2; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
423 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
424 |
structure Data = Generic_Data |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
425 |
( |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
426 |
type T = BNF Symtab.table; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
427 |
val empty = Symtab.empty; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
428 |
val extend = I; |
| 49462 | 429 |
val merge = Symtab.merge eq_bnf; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
430 |
); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
431 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
432 |
val bnf_of = Symtab.lookup o Data.get o Context.Proof; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
433 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
434 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
435 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
436 |
(* Utilities *) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
437 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
438 |
fun normalize_set insts instA set = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
439 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
440 |
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
|
441 |
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
|
442 |
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
|
443 |
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
|
444 |
|
| 49507 | 445 |
fun normalize_rel ctxt instTs instA instB rel = |
| 49462 | 446 |
let |
447 |
val thy = Proof_Context.theory_of ctxt; |
|
448 |
val tyenv = |
|
| 49507 | 449 |
Sign.typ_match thy (fastype_of rel, Library.foldr (op -->) (instTs, mk_pred2T instA instB)) |
| 49463 | 450 |
Vartab.empty; |
| 49507 | 451 |
in Envir.subst_term (tyenv, Vartab.empty) rel end |
| 49462 | 452 |
handle Type.TYPE_MATCH => error "Bad predicator"; |
453 |
||
| 49506 | 454 |
fun normalize_srel ctxt instTs instA instB srel = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
455 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
456 |
val thy = Proof_Context.theory_of ctxt; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
457 |
val tyenv = |
| 49506 | 458 |
Sign.typ_match thy (fastype_of srel, Library.foldr (op -->) (instTs, mk_relT (instA, instB))) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
459 |
Vartab.empty; |
| 49506 | 460 |
in Envir.subst_term (tyenv, Vartab.empty) srel end |
| 49453 | 461 |
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
|
462 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
463 |
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
|
464 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
465 |
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
|
466 |
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
|
467 |
| strip_param x = x; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
468 |
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
|
469 |
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
|
470 |
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
|
471 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
472 |
(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
|
473 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
474 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
475 |
fun minimize_wits wits = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
476 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
477 |
fun minimize done [] = done |
| 49103 | 478 |
| 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
|
479 |
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
|
480 |
then minimize done todo |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
481 |
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
|
482 |
in minimize [] wits end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
483 |
|
|
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 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
486 |
(* Names *) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
487 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
488 |
val mapN = "map"; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
489 |
val setN = "set"; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
490 |
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
|
491 |
val bdN = "bd"; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
492 |
val witN = "wit"; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
493 |
fun mk_witN i = witN ^ nonzero_string_of_int i; |
| 49507 | 494 |
val relN = "rel"; |
| 49506 | 495 |
val srelN = "srel"; |
|
48975
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 |
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
|
498 |
val bd_cinfiniteN = "bd_cinfinite"; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
499 |
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
|
500 |
val bd_CinfiniteN = "bd_Cinfinite"; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
501 |
val bd_CnotzeroN = "bd_Cnotzero"; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
502 |
val collect_set_naturalN = "collect_set_natural"; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
503 |
val in_bdN = "in_bd"; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
504 |
val in_monoN = "in_mono"; |
| 49506 | 505 |
val in_srelN = "in_srel"; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
506 |
val map_idN = "map_id"; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
507 |
val map_id'N = "map_id'"; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
508 |
val map_compN = "map_comp"; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
509 |
val map_comp'N = "map_comp'"; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
510 |
val map_congN = "map_cong"; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
511 |
val map_wpullN = "map_wpull"; |
|
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49585
diff
changeset
|
512 |
val rel_eqN = "rel_eq"; |
|
49537
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
513 |
val rel_flipN = "rel_flip"; |
|
49594
55e798614c45
tweaked theorem names (in particular, dropped s's)
blanchet
parents:
49592
diff
changeset
|
514 |
val rel_srelN = "rel_srel"; |
|
49537
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
515 |
val set_naturalN = "set_natural"; |
|
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
516 |
val set_natural'N = "set_natural'"; |
|
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
517 |
val set_bdN = "set_bd"; |
| 49506 | 518 |
val srel_IdN = "srel_Id"; |
519 |
val srel_GrN = "srel_Gr"; |
|
520 |
val srel_converseN = "srel_converse"; |
|
521 |
val srel_monoN = "srel_mono" |
|
522 |
val srel_ON = "srel_comp"; |
|
523 |
val srel_O_GrN = "srel_comp_Gr"; |
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
524 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
525 |
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
|
526 |
|
| 49538 | 527 |
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
|
528 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
529 |
val bnf_note_all = Attrib.setup_config_bool @{binding bnf_note_all} (K false);
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
530 |
|
| 49538 | 531 |
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
|
532 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
533 |
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
|
534 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
535 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
536 |
(* Define new BNFs *) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
537 |
|
| 49019 | 538 |
fun prepare_def const_policy mk_fact_policy qualify prep_term Ds_opt |
| 49507 | 539 |
(((((raw_b, raw_map), raw_sets), raw_bd_Abs), raw_wits), raw_rel_opt) no_defs_lthy = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
540 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
541 |
val fact_policy = mk_fact_policy no_defs_lthy; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
542 |
val b = qualify raw_b; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
543 |
val live = length raw_sets; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
544 |
val nwits = length raw_wits; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
545 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
546 |
val map_rhs = prep_term no_defs_lthy raw_map; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
547 |
val set_rhss = map (prep_term no_defs_lthy) raw_sets; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
548 |
val (bd_rhsT, bd_rhs) = (case prep_term no_defs_lthy raw_bd_Abs of |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
549 |
Abs (_, T, t) => (T, t) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
550 |
| _ => error "Bad bound constant"); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
551 |
val wit_rhss = map (prep_term no_defs_lthy) raw_wits; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
552 |
|
|
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
|
553 |
fun err T = |
|
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
|
554 |
error ("Trying to register the type " ^ quote (Syntax.string_of_typ no_defs_lthy T) ^
|
|
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
|
555 |
" as unnamed BNF"); |
|
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
|
556 |
|
|
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
|
557 |
val (b, key) = |
|
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
|
558 |
if Binding.eq_name (b, Binding.empty) then |
|
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
|
559 |
(case bd_rhsT of |
|
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
|
560 |
Type (C, Ts) => if forall (is_some o try dest_TFree) Ts |
|
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
|
561 |
then (Binding.qualified_name C, C) else err bd_rhsT |
|
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
|
562 |
| T => err T) |
|
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
|
563 |
else (b, Local_Theory.full_name no_defs_lthy b); |
|
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
|
564 |
|
| 49463 | 565 |
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
|
566 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
567 |
val inline = |
| 49538 | 568 |
(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
|
569 |
(case const_policy of |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
570 |
Dont_Inline => false |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
571 |
| 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
|
572 |
| 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
|
573 |
| Do_Inline => true) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
574 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
575 |
if inline then |
|
49460
4dd451a075ce
fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents:
49459
diff
changeset
|
576 |
((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
|
577 |
else |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
578 |
let val b = b () in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
579 |
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
|
580 |
lthy) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
581 |
end |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
582 |
end; |
| 49459 | 583 |
|
584 |
fun maybe_restore lthy_old lthy = |
|
585 |
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
|
586 |
|
| 49459 | 587 |
val map_bind_def = (fn () => Binding.suffix_name ("_" ^ mapN) b, map_rhs);
|
588 |
val set_binds_defs = |
|
589 |
let |
|
590 |
val bs = if live = 1 then [fn () => Binding.suffix_name ("_" ^ setN) b]
|
|
591 |
else map (fn i => fn () => Binding.suffix_name ("_" ^ mk_setN i) b) (1 upto live)
|
|
592 |
in map2 pair bs set_rhss end; |
|
593 |
val bd_bind_def = (fn () => Binding.suffix_name ("_" ^ bdN) b, bd_rhs);
|
|
594 |
val wit_binds_defs = |
|
595 |
let |
|
596 |
val bs = if nwits = 1 then [fn () => Binding.suffix_name ("_" ^ witN) b]
|
|
597 |
else map (fn i => fn () => Binding.suffix_name ("_" ^ mk_witN i) b) (1 upto nwits);
|
|
598 |
in map2 pair bs wit_rhss end; |
|
599 |
||
600 |
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
|
601 |
(bnf_set_terms, raw_set_defs)), |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
602 |
(bnf_bd_term, raw_bd_def)), |
| 49459 | 603 |
(bnf_wit_terms, raw_wit_defs)), (lthy, lthy_old)) = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
604 |
no_defs_lthy |
| 49463 | 605 |
|> maybe_define true map_bind_def |
606 |
||>> apfst split_list o fold_map (maybe_define true) set_binds_defs |
|
607 |
||>> maybe_define true bd_bind_def |
|
608 |
||>> apfst split_list o fold_map (maybe_define true) wit_binds_defs |
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
609 |
||> `(maybe_restore no_defs_lthy); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
610 |
|
| 49459 | 611 |
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
|
612 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
613 |
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
|
614 |
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
|
615 |
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
|
616 |
val bnf_wit_defs = map (Morphism.thm phi) raw_wit_defs; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
617 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
618 |
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
|
619 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
620 |
(*TODO: handle errors*) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
621 |
(*simple shape analysis of a map function*) |
|
49395
323414474c1f
use strip_typeN in bnf_def (instead of repairing strip_type)
traytel
parents:
49386
diff
changeset
|
622 |
val ((alphas, betas), (CA, _)) = |
|
323414474c1f
use strip_typeN in bnf_def (instead of repairing strip_type)
traytel
parents:
49386
diff
changeset
|
623 |
fastype_of bnf_map |
|
323414474c1f
use strip_typeN in bnf_def (instead of repairing strip_type)
traytel
parents:
49386
diff
changeset
|
624 |
|> strip_typeN live |
|
323414474c1f
use strip_typeN in bnf_def (instead of repairing strip_type)
traytel
parents:
49386
diff
changeset
|
625 |
|>> map_split dest_funT |
|
323414474c1f
use strip_typeN in bnf_def (instead of repairing strip_type)
traytel
parents:
49386
diff
changeset
|
626 |
||> dest_funT |
|
323414474c1f
use strip_typeN in bnf_def (instead of repairing strip_type)
traytel
parents:
49386
diff
changeset
|
627 |
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
|
628 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
629 |
val CA_params = map TVar (Term.add_tvarsT CA []); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
630 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
631 |
val bnf_sets = map2 (normalize_set CA_params) alphas (map (Morphism.term phi) bnf_set_terms); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
632 |
val bdT = Morphism.typ phi bd_rhsT; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
633 |
val bnf_bd = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
634 |
Term.subst_TVars (Term.add_tvar_namesT bdT [] ~~ CA_params) (Morphism.term phi bnf_bd_term); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
635 |
val bnf_wits = map (normalize_wit CA_params CA alphas o Morphism.term phi) bnf_wit_terms; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
636 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
637 |
(*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
|
638 |
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
|
639 |
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
|
640 |
| 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
|
641 |
val dead = length deads; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
642 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
643 |
(*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
|
644 |
(*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
|
645 |
(*TODO: check type of bnf_bd*) |
| 49507 | 646 |
(*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
|
647 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
648 |
val ((((((((((As', Bs'), Cs), Ds), B1Ts), B2Ts), domTs), ranTs), ranTs'), ranTs''), |
| 49459 | 649 |
(Ts, T)) = lthy |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
650 |
|> mk_TFrees live |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
651 |
||>> mk_TFrees live |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
652 |
||>> mk_TFrees live |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
653 |
||>> mk_TFrees dead |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
654 |
||>> mk_TFrees live |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
655 |
||>> mk_TFrees live |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
656 |
||>> mk_TFrees live |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
657 |
||>> mk_TFrees live |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
658 |
||>> mk_TFrees live |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
659 |
||>> mk_TFrees live |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
660 |
||> fst o mk_TFrees 1 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
661 |
||> the_single |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
662 |
||> `(replicate live); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
663 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
664 |
fun mk_bnf_map As' Bs' = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
665 |
Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As') @ (betas ~~ Bs')) bnf_map; |
| 49453 | 666 |
fun mk_bnf_t As' = Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As')); |
667 |
fun mk_bnf_T As' = Term.typ_subst_atomic ((deads ~~ Ds) @ (alphas ~~ As')); |
|
668 |
||
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
669 |
val (setRTs, RTs) = map_split (`HOLogic.mk_setT o HOLogic.mk_prodT) (As' ~~ Bs'); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
670 |
val setRTsAsCs = map (HOLogic.mk_setT o HOLogic.mk_prodT) (As' ~~ Cs); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
671 |
val setRTsBsCs = map (HOLogic.mk_setT o HOLogic.mk_prodT) (Bs' ~~ Cs); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
672 |
val setRT's = map (HOLogic.mk_setT o HOLogic.mk_prodT) (Bs' ~~ As'); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
673 |
val self_setRTs = map (HOLogic.mk_setT o HOLogic.mk_prodT) (As' ~~ As'); |
| 49463 | 674 |
val QTs = map2 mk_pred2T As' Bs'; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
675 |
|
| 49453 | 676 |
val CA' = mk_bnf_T As' CA; |
677 |
val CB' = mk_bnf_T Bs' CA; |
|
678 |
val CC' = mk_bnf_T Cs CA; |
|
679 |
val CRs' = mk_bnf_T RTs CA; |
|
| 49463 | 680 |
val CA'CB' = HOLogic.mk_prodT (CA', CB'); |
| 49453 | 681 |
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
682 |
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
|
683 |
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
|
684 |
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
|
685 |
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
|
686 |
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
|
687 |
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
|
688 |
val bnf_bd_As = mk_bnf_t As' bnf_bd; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
689 |
val bnf_wit_As = map (apsnd (mk_bnf_t As')) bnf_wits; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
690 |
|
|
49595
e8c57e59cbf8
got rid of other instance of shaky "Thm.generalize"
blanchet
parents:
49594
diff
changeset
|
691 |
val pre_names_lthy = lthy; |
| 49463 | 692 |
val (((((((((((((((((((((((((fs, fs_copy), gs), hs), p), (x, x')), (y, y')), (z, z')), zs), As), |
| 49459 | 693 |
As_copy), Xs), B1s), B2s), f1s), f2s), e1s), e2s), p1s), p2s), bs), (Rs, Rs')), Rs_copy), Ss), |
|
49595
e8c57e59cbf8
got rid of other instance of shaky "Thm.generalize"
blanchet
parents:
49594
diff
changeset
|
694 |
(Qs, Qs')), names_lthy) = pre_names_lthy |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
695 |
|> mk_Frees "f" (map2 (curry (op -->)) As' Bs') |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
696 |
||>> mk_Frees "f" (map2 (curry (op -->)) As' Bs') |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
697 |
||>> mk_Frees "g" (map2 (curry (op -->)) Bs' Cs) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
698 |
||>> mk_Frees "h" (map2 (curry (op -->)) As' Ts) |
| 49463 | 699 |
||>> yield_singleton (mk_Frees "p") CA'CB' |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
700 |
||>> yield_singleton (apfst (op ~~) oo mk_Frees' "x") CA' |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
701 |
||>> yield_singleton (apfst (op ~~) oo mk_Frees' "y") CB' |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
702 |
||>> yield_singleton (apfst (op ~~) oo mk_Frees' "z") CRs' |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
703 |
||>> mk_Frees "z" As' |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
704 |
||>> 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
|
705 |
||>> 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
|
706 |
||>> mk_Frees "A" (map HOLogic.mk_setT domTs) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
707 |
||>> mk_Frees "B1" (map HOLogic.mk_setT B1Ts) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
708 |
||>> mk_Frees "B2" (map HOLogic.mk_setT B2Ts) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
709 |
||>> mk_Frees "f1" (map2 (curry (op -->)) B1Ts ranTs) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
710 |
||>> mk_Frees "f2" (map2 (curry (op -->)) B2Ts ranTs) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
711 |
||>> mk_Frees "e1" (map2 (curry (op -->)) B1Ts ranTs') |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
712 |
||>> mk_Frees "e2" (map2 (curry (op -->)) B2Ts ranTs'') |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
713 |
||>> mk_Frees "p1" (map2 (curry (op -->)) domTs B1Ts) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
714 |
||>> mk_Frees "p2" (map2 (curry (op -->)) domTs B2Ts) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
715 |
||>> mk_Frees "b" As' |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49538
diff
changeset
|
716 |
||>> mk_Frees' "r" setRTs |
|
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49538
diff
changeset
|
717 |
||>> mk_Frees "r" setRTs |
|
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49538
diff
changeset
|
718 |
||>> mk_Frees "s" setRTsBsCs |
|
49592
b859a02c1150
fixed "rels" + split them into injectivity and distinctness
blanchet
parents:
49591
diff
changeset
|
719 |
||>> mk_Frees' "P" QTs; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
720 |
|
| 49459 | 721 |
(*Gr (in R1 .. Rn) (map fst .. fst)^-1 O Gr (in R1 .. Rn) (map snd .. snd)*) |
| 49463 | 722 |
val O_Gr = |
| 49459 | 723 |
let |
724 |
val map1 = Term.list_comb (mk_bnf_map RTs As', map fst_const RTs); |
|
725 |
val map2 = Term.list_comb (mk_bnf_map RTs Bs', map snd_const RTs); |
|
726 |
val bnf_in = mk_in (map Free Rs') (map (mk_bnf_t RTs) bnf_sets) CRs'; |
|
727 |
in |
|
728 |
mk_rel_comp (mk_converse (mk_Gr bnf_in map1), mk_Gr bnf_in map2) |
|
729 |
end; |
|
730 |
||
| 49463 | 731 |
fun mk_predicate_of_set x_name y_name t = |
732 |
let |
|
733 |
val (T, U) = HOLogic.dest_prodT (HOLogic.dest_setT (fastype_of t)); |
|
734 |
val x = Free (x_name, T); |
|
735 |
val y = Free (y_name, U); |
|
736 |
in fold_rev Term.lambda [x, y] (HOLogic.mk_mem (HOLogic.mk_prod (x, y), t)) end; |
|
737 |
||
|
49537
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
738 |
val sQs = |
|
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
739 |
map3 (fn Q => fn T => fn U => |
|
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
740 |
HOLogic.Collect_const (HOLogic.mk_prodT (T, U)) $ HOLogic.mk_split Q) Qs As' Bs'; |
|
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
741 |
|
| 49507 | 742 |
val rel_rhs = (case raw_rel_opt of |
| 49463 | 743 |
NONE => |
744 |
fold_rev absfree Qs' (mk_predicate_of_set (fst x') (fst y') |
|
|
49537
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
745 |
(Term.list_comb (fold_rev Term.absfree Rs' O_Gr, sQs))) |
| 49507 | 746 |
| SOME raw_rel => prep_term no_defs_lthy raw_rel); |
| 49463 | 747 |
|
| 49507 | 748 |
val rel_bind_def = (fn () => Binding.suffix_name ("_" ^ relN) b, rel_rhs);
|
| 49463 | 749 |
|
| 49507 | 750 |
val ((bnf_rel_term, raw_rel_def), (lthy, lthy_old)) = |
| 49463 | 751 |
lthy |
| 49507 | 752 |
|> maybe_define (is_some raw_rel_opt) rel_bind_def |
| 49463 | 753 |
||> `(maybe_restore lthy); |
754 |
||
755 |
val phi = Proof_Context.export_morphism lthy_old lthy; |
|
| 49507 | 756 |
val bnf_rel_def = Morphism.thm phi raw_rel_def; |
757 |
val bnf_rel = Morphism.term phi bnf_rel_term; |
|
| 49463 | 758 |
|
| 49507 | 759 |
fun mk_bnf_rel QTs CA' CB' = normalize_rel lthy QTs CA' CB' bnf_rel; |
| 49463 | 760 |
|
| 49507 | 761 |
val rel = mk_bnf_rel QTs CA' CB'; |
| 49463 | 762 |
|
| 49506 | 763 |
val srel_rhs = |
| 49463 | 764 |
fold_rev Term.absfree Rs' (HOLogic.Collect_const CA'CB' $ |
| 49507 | 765 |
Term.lambda p (Term.list_comb (rel, map (mk_predicate_of_set (fst x') (fst y')) Rs) $ |
| 49463 | 766 |
HOLogic.mk_fst p $ HOLogic.mk_snd p)); |
| 49459 | 767 |
|
| 49506 | 768 |
val srel_bind_def = (fn () => Binding.suffix_name ("_" ^ srelN) b, srel_rhs);
|
| 49459 | 769 |
|
| 49506 | 770 |
val ((bnf_srel_term, raw_srel_def), (lthy, lthy_old)) = |
| 49459 | 771 |
lthy |
| 49506 | 772 |
|> maybe_define false srel_bind_def |
| 49459 | 773 |
||> `(maybe_restore lthy); |
774 |
||
775 |
val phi = Proof_Context.export_morphism lthy_old lthy; |
|
| 49506 | 776 |
val bnf_srel_def = Morphism.thm phi raw_srel_def; |
777 |
val bnf_srel = Morphism.term phi bnf_srel_term; |
|
| 49459 | 778 |
|
| 49506 | 779 |
fun mk_bnf_srel setRTs CA' CB' = normalize_srel lthy setRTs CA' CB' bnf_srel; |
| 49459 | 780 |
|
| 49506 | 781 |
val srel = mk_bnf_srel setRTs CA' CB'; |
| 49459 | 782 |
|
|
49460
4dd451a075ce
fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents:
49459
diff
changeset
|
783 |
val _ = case no_reflexive (raw_map_def :: raw_set_defs @ [raw_bd_def] @ |
| 49507 | 784 |
raw_wit_defs @ [raw_rel_def, raw_srel_def]) of |
| 49459 | 785 |
[] => () |
786 |
| defs => Proof_Display.print_consts true lthy_old (K false) |
|
787 |
(map (dest_Free o fst o Logic.dest_equals o prop_of) defs); |
|
788 |
||
| 49458 | 789 |
val map_id_goal = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
790 |
let |
| 49018 | 791 |
val bnf_map_app_id = Term.list_comb (bnf_map_AsAs, map HOLogic.id_const As'); |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
792 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
793 |
HOLogic.mk_Trueprop |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
794 |
(HOLogic.mk_eq (bnf_map_app_id, HOLogic.id_const CA')) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
795 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
796 |
|
| 49458 | 797 |
val map_comp_goal = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
798 |
let |
| 49018 | 799 |
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
|
800 |
val comp_bnf_map_app = HOLogic.mk_comp |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
801 |
(Term.list_comb (bnf_map_BsCs, gs), |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
802 |
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
|
803 |
in |
|
49123
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
49111
diff
changeset
|
804 |
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
|
805 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
806 |
|
| 49458 | 807 |
val map_cong_goal = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
808 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
809 |
fun mk_prem z set f f_copy = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
810 |
Logic.all z (Logic.mk_implies |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
811 |
(HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set $ x)), |
|
49123
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
49111
diff
changeset
|
812 |
mk_Trueprop_eq (f $ z, f_copy $ z))); |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
813 |
val prems = map4 mk_prem zs bnf_sets_As fs fs_copy; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
814 |
val eq = HOLogic.mk_eq (Term.list_comb (bnf_map_AsBs, fs) $ x, |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
815 |
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
|
816 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
817 |
fold_rev Logic.all (x :: fs @ fs_copy) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
818 |
(Logic.list_implies (prems, HOLogic.mk_Trueprop eq)) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
819 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
820 |
|
| 49458 | 821 |
val set_naturals_goal = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
822 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
823 |
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
|
824 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
825 |
val set_comp_map = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
826 |
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
|
827 |
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
|
828 |
in |
|
49123
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
49111
diff
changeset
|
829 |
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
|
830 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
831 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
832 |
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
|
833 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
834 |
|
| 49458 | 835 |
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
|
836 |
|
| 49458 | 837 |
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
|
838 |
|
| 49458 | 839 |
val set_bds_goal = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
840 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
841 |
fun mk_goal set = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
842 |
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
|
843 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
844 |
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
|
845 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
846 |
|
| 49458 | 847 |
val in_bd_goal = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
848 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
849 |
val bd = mk_cexp |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
850 |
(if live = 0 then ctwo |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
851 |
else mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
852 |
bnf_bd_As; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
853 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
854 |
fold_rev Logic.all As |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
855 |
(HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (mk_in As bnf_sets_As CA')) bd)) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
856 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
857 |
|
| 49458 | 858 |
val map_wpull_goal = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
859 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
860 |
val prems = map HOLogic.mk_Trueprop |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
861 |
(map8 mk_wpull Xs B1s B2s f1s f2s (replicate live NONE) p1s p2s); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
862 |
val CX = mk_bnf_T domTs CA; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
863 |
val CB1 = mk_bnf_T B1Ts CA; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
864 |
val CB2 = mk_bnf_T B2Ts CA; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
865 |
val bnf_sets_CX = map2 (normalize_set (map (mk_bnf_T domTs) CA_params)) domTs bnf_sets; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
866 |
val bnf_sets_CB1 = map2 (normalize_set (map (mk_bnf_T B1Ts) CA_params)) B1Ts bnf_sets; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
867 |
val bnf_sets_CB2 = map2 (normalize_set (map (mk_bnf_T B2Ts) CA_params)) B2Ts bnf_sets; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
868 |
val bnf_map_app_f1 = Term.list_comb (mk_bnf_map B1Ts ranTs, f1s); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
869 |
val bnf_map_app_f2 = Term.list_comb (mk_bnf_map B2Ts ranTs, f2s); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
870 |
val bnf_map_app_p1 = Term.list_comb (mk_bnf_map domTs B1Ts, p1s); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
871 |
val bnf_map_app_p2 = Term.list_comb (mk_bnf_map domTs B2Ts, p2s); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
872 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
873 |
val map_wpull = mk_wpull (mk_in Xs bnf_sets_CX CX) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
874 |
(mk_in B1s bnf_sets_CB1 CB1) (mk_in B2s bnf_sets_CB2 CB2) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
875 |
bnf_map_app_f1 bnf_map_app_f2 NONE bnf_map_app_p1 bnf_map_app_p2; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
876 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
877 |
fold_rev Logic.all (Xs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
878 |
(Logic.list_implies (prems, HOLogic.mk_Trueprop map_wpull)) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
879 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
880 |
|
| 49506 | 881 |
val srel_O_Gr_goal = fold_rev Logic.all Rs (mk_Trueprop_eq (Term.list_comb (srel, Rs), O_Gr)); |
| 49453 | 882 |
|
|
49460
4dd451a075ce
fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents:
49459
diff
changeset
|
883 |
val goals = zip_axioms map_id_goal map_comp_goal map_cong_goal set_naturals_goal |
| 49506 | 884 |
card_order_bd_goal cinfinite_bd_goal set_bds_goal in_bd_goal map_wpull_goal srel_O_Gr_goal; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
885 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
886 |
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
|
887 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
888 |
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
|
889 |
fun wit_goal i = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
890 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
891 |
val z = nth zs i; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
892 |
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
|
893 |
val concl = HOLogic.mk_Trueprop |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
894 |
(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
|
895 |
else @{term False});
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
896 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
897 |
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
|
898 |
(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
|
899 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
900 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
901 |
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
|
902 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
903 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
904 |
val wit_goalss = map mk_wit_goals bnf_wit_As; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
905 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
906 |
fun after_qed thms lthy = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
907 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
908 |
val (axioms, wit_thms) = apfst (mk_axioms live) (chop (length goals) thms); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
909 |
|
|
49109
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
910 |
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
|
911 |
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
|
912 |
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
|
913 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
914 |
fun mk_collect_set_natural () = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
915 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
916 |
val defT = mk_bnf_T Ts CA --> HOLogic.mk_setT T; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
917 |
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
|
918 |
(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
|
919 |
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
|
920 |
val image_collect = mk_collect |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
921 |
(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
|
922 |
defT; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
923 |
(*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
|
924 |
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
|
925 |
in |
| 49623 | 926 |
Skip_Proof.prove lthy [] [] goal (K (mk_collect_set_natural_tac (#set_natural axioms))) |
|
49109
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
927 |
|> Thm.close_derivation |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
928 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
929 |
|
| 49538 | 930 |
val collect_set_natural = Lazy.lazy mk_collect_set_natural; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
931 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
932 |
fun mk_in_mono () = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
933 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
934 |
val prems_mono = map2 (HOLogic.mk_Trueprop oo mk_subset) As As_copy; |
| 49458 | 935 |
val in_mono_goal = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
936 |
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
|
937 |
(Logic.list_implies (prems_mono, HOLogic.mk_Trueprop |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
938 |
(mk_subset (mk_in As bnf_sets_As CA') (mk_in As_copy bnf_sets_As CA')))); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
939 |
in |
| 49458 | 940 |
Skip_Proof.prove 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
|
941 |
|> Thm.close_derivation |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
942 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
943 |
|
| 49538 | 944 |
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
|
945 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
946 |
fun mk_in_cong () = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
947 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
948 |
val prems_cong = map2 (HOLogic.mk_Trueprop oo curry HOLogic.mk_eq) As As_copy; |
| 49458 | 949 |
val in_cong_goal = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
950 |
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
|
951 |
(Logic.list_implies (prems_cong, HOLogic.mk_Trueprop |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
952 |
(HOLogic.mk_eq (mk_in As bnf_sets_As CA', mk_in As_copy bnf_sets_As CA')))); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
953 |
in |
| 49458 | 954 |
Skip_Proof.prove lthy [] [] in_cong_goal (K ((TRY o hyp_subst_tac 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
|
955 |
|> Thm.close_derivation |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
956 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
957 |
|
| 49538 | 958 |
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
|
959 |
|
| 49538 | 960 |
val map_id' = Lazy.lazy (fn () => mk_id' (#map_id axioms)); |
961 |
val map_comp' = Lazy.lazy (fn () => mk_comp' (#map_comp axioms)); |
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
962 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
963 |
val set_natural' = |
| 49538 | 964 |
map (fn thm => Lazy.lazy (fn () => mk_set_natural' thm)) (#set_natural axioms); |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
965 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
966 |
fun mk_map_wppull () = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
967 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
968 |
val prems = if live = 0 then [] else |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
969 |
[HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
970 |
(map8 mk_wpull Xs B1s B2s f1s f2s (map SOME (e1s ~~ e2s)) p1s p2s))]; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
971 |
val CX = mk_bnf_T domTs CA; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
972 |
val CB1 = mk_bnf_T B1Ts CA; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
973 |
val CB2 = mk_bnf_T B2Ts CA; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
974 |
val bnf_sets_CX = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
975 |
map2 (normalize_set (map (mk_bnf_T domTs) CA_params)) domTs bnf_sets; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
976 |
val bnf_sets_CB1 = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
977 |
map2 (normalize_set (map (mk_bnf_T B1Ts) CA_params)) B1Ts bnf_sets; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
978 |
val bnf_sets_CB2 = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
979 |
map2 (normalize_set (map (mk_bnf_T B2Ts) CA_params)) B2Ts bnf_sets; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
980 |
val bnf_map_app_f1 = Term.list_comb (mk_bnf_map B1Ts ranTs, f1s); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
981 |
val bnf_map_app_f2 = Term.list_comb (mk_bnf_map B2Ts ranTs, f2s); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
982 |
val bnf_map_app_e1 = Term.list_comb (mk_bnf_map B1Ts ranTs', e1s); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
983 |
val bnf_map_app_e2 = Term.list_comb (mk_bnf_map B2Ts ranTs'', e2s); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
984 |
val bnf_map_app_p1 = Term.list_comb (mk_bnf_map domTs B1Ts, p1s); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
985 |
val bnf_map_app_p2 = Term.list_comb (mk_bnf_map domTs B2Ts, p2s); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
986 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
987 |
val concl = mk_wpull (mk_in Xs bnf_sets_CX CX) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
988 |
(mk_in B1s bnf_sets_CB1 CB1) (mk_in B2s bnf_sets_CB2 CB2) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
989 |
bnf_map_app_f1 bnf_map_app_f2 (SOME (bnf_map_app_e1, bnf_map_app_e2)) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
990 |
bnf_map_app_p1 bnf_map_app_p2; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
991 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
992 |
val goal = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
993 |
fold_rev Logic.all (Xs @ B1s @ B2s @ f1s @ f2s @ e1s @ e2s @ p1s @ p2s) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
994 |
(Logic.list_implies (prems, HOLogic.mk_Trueprop concl)) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
995 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
996 |
Skip_Proof.prove lthy [] [] goal |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
997 |
(fn _ => mk_map_wppull_tac (#map_id axioms) (#map_cong axioms) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
998 |
(#map_wpull axioms) (Lazy.force map_comp') (map Lazy.force set_natural')) |
|
49109
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
999 |
|> Thm.close_derivation |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1000 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1001 |
|
| 49506 | 1002 |
val srel_O_Grs = no_refl [#srel_O_Gr axioms]; |
| 49453 | 1003 |
|
| 49538 | 1004 |
val map_wppull = Lazy.lazy mk_map_wppull; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1005 |
|
| 49506 | 1006 |
fun mk_srel_Gr () = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1007 |
let |
| 49506 | 1008 |
val lhs = Term.list_comb (srel, map2 mk_Gr As fs); |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1009 |
val rhs = mk_Gr (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
|
1010 |
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
|
1011 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1012 |
Skip_Proof.prove lthy [] [] goal |
| 49506 | 1013 |
(mk_srel_Gr_tac srel_O_Grs (#map_id axioms) (#map_cong axioms) (Lazy.force map_id') |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1014 |
(Lazy.force map_comp') (map Lazy.force set_natural')) |
|
49109
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1015 |
|> Thm.close_derivation |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1016 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1017 |
|
| 49538 | 1018 |
val srel_Gr = Lazy.lazy mk_srel_Gr; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1019 |
|
| 49506 | 1020 |
fun mk_srel_prems f = map2 (HOLogic.mk_Trueprop oo f) Rs Rs_copy |
1021 |
fun mk_srel_concl f = HOLogic.mk_Trueprop |
|
1022 |
(f (Term.list_comb (srel, Rs), Term.list_comb (srel, Rs_copy))); |
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1023 |
|
| 49506 | 1024 |
fun mk_srel_mono () = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1025 |
let |
| 49506 | 1026 |
val mono_prems = mk_srel_prems mk_subset; |
1027 |
val mono_concl = mk_srel_concl (uncurry mk_subset); |
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1028 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1029 |
Skip_Proof.prove lthy [] [] |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1030 |
(fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (mono_prems, mono_concl))) |
| 49506 | 1031 |
(mk_srel_mono_tac srel_O_Grs (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
|
1032 |
|> Thm.close_derivation |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1033 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1034 |
|
| 49506 | 1035 |
fun mk_srel_cong () = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1036 |
let |
| 49506 | 1037 |
val cong_prems = mk_srel_prems (curry HOLogic.mk_eq); |
1038 |
val cong_concl = mk_srel_concl HOLogic.mk_eq; |
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1039 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1040 |
Skip_Proof.prove lthy [] [] |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1041 |
(fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (cong_prems, cong_concl))) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1042 |
(fn _ => (TRY o hyp_subst_tac 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
|
1043 |
|> Thm.close_derivation |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1044 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1045 |
|
| 49538 | 1046 |
val srel_mono = Lazy.lazy mk_srel_mono; |
1047 |
val srel_cong = Lazy.lazy mk_srel_cong; |
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1048 |
|
| 49506 | 1049 |
fun mk_srel_Id () = |
| 49515 | 1050 |
let val srelAsAs = mk_bnf_srel self_setRTs CA' CA' in |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1051 |
Skip_Proof.prove lthy [] [] |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1052 |
(HOLogic.mk_Trueprop |
| 49515 | 1053 |
(HOLogic.mk_eq (Term.list_comb (srelAsAs, map Id_const As'), Id_const CA'))) |
| 49506 | 1054 |
(mk_srel_Id_tac live (Lazy.force srel_Gr) (#map_id axioms)) |
|
49109
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1055 |
|> Thm.close_derivation |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1056 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1057 |
|
| 49538 | 1058 |
val srel_Id = Lazy.lazy mk_srel_Id; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1059 |
|
| 49506 | 1060 |
fun mk_srel_converse () = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1061 |
let |
| 49515 | 1062 |
val srelBsAs = mk_bnf_srel setRT's CB' CA'; |
1063 |
val lhs = Term.list_comb (srelBsAs, map mk_converse Rs); |
|
| 49506 | 1064 |
val rhs = mk_converse (Term.list_comb (srel, Rs)); |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1065 |
val le_goal = fold_rev Logic.all Rs (HOLogic.mk_Trueprop (mk_subset lhs rhs)); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1066 |
val le_thm = Skip_Proof.prove lthy [] [] le_goal |
| 49506 | 1067 |
(mk_srel_converse_le_tac srel_O_Grs (Lazy.force srel_Id) (#map_cong axioms) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1068 |
(Lazy.force map_comp') (map Lazy.force set_natural')) |
|
49109
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1069 |
|> Thm.close_derivation |
|
49123
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
49111
diff
changeset
|
1070 |
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
|
1071 |
in |
| 49506 | 1072 |
Skip_Proof.prove lthy [] [] goal (fn _ => mk_srel_converse_tac le_thm) |
|
49109
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1073 |
|> Thm.close_derivation |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1074 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1075 |
|
| 49538 | 1076 |
val srel_converse = Lazy.lazy mk_srel_converse; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1077 |
|
| 49506 | 1078 |
fun mk_srel_O () = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1079 |
let |
| 49515 | 1080 |
val srelAsCs = mk_bnf_srel setRTsAsCs CA' CC'; |
1081 |
val srelBsCs = mk_bnf_srel setRTsBsCs CB' CC'; |
|
1082 |
val lhs = Term.list_comb (srelAsCs, map2 (curry mk_rel_comp) Rs Ss); |
|
1083 |
val rhs = mk_rel_comp (Term.list_comb (srel, Rs), Term.list_comb (srelBsCs, Ss)); |
|
|
49123
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
49111
diff
changeset
|
1084 |
val goal = fold_rev Logic.all (Rs @ Ss) (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
|
1085 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1086 |
Skip_Proof.prove lthy [] [] goal |
| 49506 | 1087 |
(mk_srel_O_tac srel_O_Grs (Lazy.force srel_Id) (#map_cong axioms) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1088 |
(Lazy.force map_wppull) (Lazy.force map_comp') (map Lazy.force set_natural')) |
|
49109
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1089 |
|> Thm.close_derivation |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1090 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1091 |
|
| 49538 | 1092 |
val srel_O = Lazy.lazy mk_srel_O; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1093 |
|
| 49506 | 1094 |
fun mk_in_srel () = |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1095 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1096 |
val bnf_in = mk_in Rs (map (mk_bnf_t RTs) bnf_sets) CRs'; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1097 |
val map1 = Term.list_comb (mk_bnf_map RTs As', map fst_const RTs); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1098 |
val map2 = Term.list_comb (mk_bnf_map RTs Bs', map snd_const RTs); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1099 |
val map_fst_eq = HOLogic.mk_eq (map1 $ z, x); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1100 |
val map_snd_eq = HOLogic.mk_eq (map2 $ z, y); |
| 49506 | 1101 |
val lhs = HOLogic.mk_mem (HOLogic.mk_prod (x, y), Term.list_comb (srel, Rs)); |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1102 |
val rhs = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1103 |
HOLogic.mk_exists (fst z', snd z', HOLogic.mk_conj (HOLogic.mk_mem (z, bnf_in), |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1104 |
HOLogic.mk_conj (map_fst_eq, map_snd_eq))); |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1105 |
val goal = |
|
49123
263b0e330d8b
more work on sugar + simplify Trueprop + eq idiom everywhere
blanchet
parents:
49111
diff
changeset
|
1106 |
fold_rev Logic.all (x :: y :: 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
|
1107 |
in |
| 49506 | 1108 |
Skip_Proof.prove lthy [] [] goal (mk_in_srel_tac srel_O_Grs (length bnf_sets)) |
|
49109
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1109 |
|> Thm.close_derivation |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1110 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1111 |
|
| 49538 | 1112 |
val in_srel = Lazy.lazy mk_in_srel; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1113 |
|
|
49537
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
1114 |
val eqset_imp_iff_pair = @{thm eqset_imp_iff_pair};
|
|
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49585
diff
changeset
|
1115 |
val mem_Collect_etc = @{thms fst_conv mem_Collect_eq prod.cases snd_conv};
|
|
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49585
diff
changeset
|
1116 |
val mem_Collect_etc' = @{thms fst_conv mem_Collect_eq pair_in_Id_conv snd_conv};
|
|
49537
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
1117 |
|
|
49594
55e798614c45
tweaked theorem names (in particular, dropped s's)
blanchet
parents:
49592
diff
changeset
|
1118 |
fun mk_rel_srel () = |
|
49537
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
1119 |
unfold_thms lthy mem_Collect_etc |
|
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
1120 |
(funpow live (fn thm => thm RS @{thm fun_cong_pair}) (bnf_srel_def RS meta_eq_to_obj_eq)
|
|
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
1121 |
RS eqset_imp_iff_pair RS sym) |
|
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
1122 |
|> Drule.zero_var_indexes; |
|
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
1123 |
|
|
49594
55e798614c45
tweaked theorem names (in particular, dropped s's)
blanchet
parents:
49592
diff
changeset
|
1124 |
val rel_srel = Lazy.lazy mk_rel_srel; |
|
49537
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
1125 |
|
|
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49585
diff
changeset
|
1126 |
fun mk_rel_eq () = |
|
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49585
diff
changeset
|
1127 |
unfold_thms lthy (bnf_srel_def :: mem_Collect_etc') |
| 49631 | 1128 |
(Lazy.force srel_Id RS @{thm arg_cong[of _ _ "%A x y. (x, y) : A"]})
|
1129 |
|> Drule.eta_contraction_rule; |
|
|
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49585
diff
changeset
|
1130 |
|
|
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49585
diff
changeset
|
1131 |
val rel_eq = Lazy.lazy mk_rel_eq; |
|
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49585
diff
changeset
|
1132 |
|
|
49537
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
1133 |
fun mk_rel_flip () = |
|
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
1134 |
let |
|
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
1135 |
val srel_converse_thm = Lazy.force srel_converse; |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49538
diff
changeset
|
1136 |
val cts = map (SOME o certify lthy) sQs; |
|
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49538
diff
changeset
|
1137 |
val srel_converse_thm' = cterm_instantiate_pos cts srel_converse_thm; |
|
49537
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
1138 |
in |
|
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
1139 |
unfold_thms lthy (bnf_srel_def :: @{thm converse_iff} :: mem_Collect_etc)
|
|
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
1140 |
(srel_converse_thm' RS eqset_imp_iff_pair) |
|
49595
e8c57e59cbf8
got rid of other instance of shaky "Thm.generalize"
blanchet
parents:
49594
diff
changeset
|
1141 |
|> singleton (Proof_Context.export names_lthy pre_names_lthy) |
|
49537
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
1142 |
end; |
|
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
1143 |
|
| 49538 | 1144 |
val rel_flip = Lazy.lazy mk_rel_flip; |
|
49537
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
1145 |
|
| 49507 | 1146 |
val defs = mk_defs bnf_map_def bnf_set_defs bnf_rel_def bnf_srel_def; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1147 |
|
| 49506 | 1148 |
val facts = mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_natural in_cong |
|
49594
55e798614c45
tweaked theorem names (in particular, dropped s's)
blanchet
parents:
49592
diff
changeset
|
1149 |
in_mono in_srel map_comp' map_id' map_wppull rel_eq rel_flip rel_srel set_natural' |
|
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49585
diff
changeset
|
1150 |
srel_cong srel_mono srel_Id srel_Gr srel_converse srel_O; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1151 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1152 |
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
|
1153 |
|
| 49507 | 1154 |
val bnf_rel = |
1155 |
Term.subst_atomic_types ((Ds ~~ deads) @ (As' ~~ alphas) @ (Bs' ~~ betas)) rel; |
|
| 49506 | 1156 |
val bnf_srel = |
1157 |
Term.subst_atomic_types ((Ds ~~ deads) @ (As' ~~ alphas) @ (Bs' ~~ betas)) srel; |
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1158 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1159 |
val bnf = mk_bnf b CA live alphas betas dead deads bnf_map bnf_sets bnf_bd axioms defs facts |
| 49507 | 1160 |
wits bnf_rel bnf_srel; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1161 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1162 |
(bnf, lthy |
| 49538 | 1163 |
|> (if fact_policy = Note_All then |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1164 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1165 |
val witNs = if length wits = 1 then [witN] else map mk_witN (1 upto length wits); |
|
49109
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1166 |
val notes = |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1167 |
[(bd_card_orderN, [#bd_card_order axioms]), |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1168 |
(bd_cinfiniteN, [#bd_cinfinite axioms]), |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1169 |
(bd_Card_orderN, [#bd_Card_order facts]), |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1170 |
(bd_CinfiniteN, [#bd_Cinfinite facts]), |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1171 |
(bd_CnotzeroN, [#bd_Cnotzero facts]), |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1172 |
(collect_set_naturalN, [Lazy.force (#collect_set_natural facts)]), |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1173 |
(in_bdN, [#in_bd axioms]), |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1174 |
(in_monoN, [Lazy.force (#in_mono facts)]), |
| 49506 | 1175 |
(in_srelN, [Lazy.force (#in_srel facts)]), |
|
49109
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1176 |
(map_compN, [#map_comp axioms]), |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1177 |
(map_idN, [#map_id axioms]), |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1178 |
(map_wpullN, [#map_wpull axioms]), |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1179 |
(set_naturalN, #set_natural axioms), |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1180 |
(set_bdN, #set_bd axioms)] @ |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1181 |
map2 pair witNs wit_thms |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1182 |
|> map (fn (thmN, thms) => |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1183 |
((qualify (Binding.qualify true (Binding.name_of b) (Binding.name thmN)), []), |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1184 |
[(thms, [])])); |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1185 |
in |
|
49109
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1186 |
Local_Theory.notes notes #> snd |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1187 |
end |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1188 |
else |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1189 |
I) |
| 49538 | 1190 |
|> (if fact_policy <> Dont_Note then |
|
49109
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1191 |
let |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1192 |
val notes = |
| 49506 | 1193 |
[(map_comp'N, [Lazy.force (#map_comp' facts)]), |
1194 |
(map_congN, [#map_cong axioms]), |
|
|
49109
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1195 |
(map_id'N, [Lazy.force (#map_id' facts)]), |
|
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49585
diff
changeset
|
1196 |
(rel_eqN, [Lazy.force (#rel_eq facts)]), |
|
49537
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
blanchet
parents:
49536
diff
changeset
|
1197 |
(rel_flipN, [Lazy.force (#rel_flip facts)]), |
|
49594
55e798614c45
tweaked theorem names (in particular, dropped s's)
blanchet
parents:
49592
diff
changeset
|
1198 |
(rel_srelN, [Lazy.force (#rel_srel facts)]), |
| 49506 | 1199 |
(set_natural'N, map Lazy.force (#set_natural' facts)), |
1200 |
(srel_O_GrN, srel_O_Grs), |
|
1201 |
(srel_IdN, [Lazy.force (#srel_Id facts)]), |
|
1202 |
(srel_GrN, [Lazy.force (#srel_Gr facts)]), |
|
1203 |
(srel_converseN, [Lazy.force (#srel_converse facts)]), |
|
1204 |
(srel_monoN, [Lazy.force (#srel_mono facts)]), |
|
1205 |
(srel_ON, [Lazy.force (#srel_O facts)])] |
|
|
49460
4dd451a075ce
fixed infinite loop with trivial rel_O_Gr + tuning
blanchet
parents:
49459
diff
changeset
|
1206 |
|> filter_out (null o #2) |
|
49109
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1207 |
|> map (fn (thmN, thms) => |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1208 |
((qualify (Binding.qualify true (Binding.name_of b) (Binding.name thmN)), []), |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1209 |
[(thms, [])])); |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1210 |
in |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1211 |
Local_Theory.notes notes #> snd |
|
0e5b859e1c91
no more aliases for Local_Theory.note; use Thm.close_derivation in internal theorems;
traytel
parents:
49103
diff
changeset
|
1212 |
end |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1213 |
else |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1214 |
I)) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1215 |
end; |
| 49459 | 1216 |
|
1217 |
val one_step_defs = |
|
| 49507 | 1218 |
no_reflexive (bnf_map_def :: bnf_bd_def :: bnf_set_defs @ bnf_wit_defs @ [bnf_rel_def, |
| 49506 | 1219 |
bnf_srel_def]); |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1220 |
in |
| 49459 | 1221 |
(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
|
1222 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1223 |
|
|
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
|
1224 |
fun register_bnf key (bnf, lthy) = |
|
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
|
1225 |
(bnf, Local_Theory.declaration {syntax = false, pervasive = true}
|
|
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
|
1226 |
(fn phi => Data.map (Symtab.update_new (key, morph_bnf phi bnf))) lthy); |
|
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
|
1227 |
|
| 49456 | 1228 |
(* TODO: Once the invariant "nwits > 0" holds, remove "mk_conjunction_balanced'" and "rtac TrueI" |
1229 |
below *) |
|
1230 |
fun mk_conjunction_balanced' [] = @{prop True}
|
|
1231 |
| mk_conjunction_balanced' ts = Logic.mk_conjunction_balanced ts; |
|
1232 |
||
| 49018 | 1233 |
fun bnf_def const_policy fact_policy qualify tacs wit_tac Ds = |
| 49463 | 1234 |
(fn (_, 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
|
1235 |
let |
| 49456 | 1236 |
val wits_tac = |
1237 |
K (TRYALL Goal.conjunction_tac) THEN' K (TRYALL (rtac TrueI)) THEN' |
|
|
49504
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents:
49495
diff
changeset
|
1238 |
mk_unfold_thms_then_tac lthy one_step_defs wit_tac; |
| 49456 | 1239 |
val wit_goals = map mk_conjunction_balanced' wit_goalss; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1240 |
val wit_thms = |
| 49456 | 1241 |
Skip_Proof.prove lthy [] [] (mk_conjunction_balanced' wit_goals) wits_tac |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1242 |
|> Conjunction.elim_balanced (length wit_goals) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1243 |
|> map2 (Conjunction.elim_balanced o length) wit_goalss |
| 49456 | 1244 |
|> 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
|
1245 |
in |
| 49111 | 1246 |
map2 (Thm.close_derivation oo Skip_Proof.prove lthy [] []) |
|
49504
df9b897fb254
renamed "iter"/"coiter" to "fold"/"unfold" (cf. Wadler)
blanchet
parents:
49495
diff
changeset
|
1247 |
goals (map (mk_unfold_thms_then_tac lthy one_step_defs) tacs) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1248 |
|> (fn thms => after_qed (map single thms @ wit_thms) lthy) |
| 49279 | 1249 |
end) oo prepare_def const_policy fact_policy qualify (K I) Ds; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1250 |
|
|
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
|
1251 |
val bnf_def_cmd = (fn (key, goals, wit_goals, after_qed, lthy, defs) => |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1252 |
Proof.unfolding ([[(defs, [])]]) |
|
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
|
1253 |
(Proof.theorem NONE (snd o register_bnf key oo after_qed) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1254 |
(map (single o rpair []) goals @ map (map (rpair [])) wit_goals) lthy)) oo |
| 49538 | 1255 |
prepare_def Do_Inline (user_policy Note_Some) I Syntax.read_term NONE; |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1256 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1257 |
fun print_bnfs ctxt = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1258 |
let |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1259 |
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
|
1260 |
[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
|
1261 |
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
|
1262 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1263 |
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
|
1264 |
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
|
1265 |
Pretty.big_list |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1266 |
(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
|
1267 |
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
|
1268 |
([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
|
1269 |
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
|
1270 |
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
|
1271 |
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
|
1272 |
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
|
1273 |
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
|
1274 |
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
|
1275 |
[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
|
1276 |
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
|
1277 |
in |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1278 |
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
|
1279 |
|> Pretty.writeln |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1280 |
end; |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1281 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1282 |
val _ = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1283 |
Outer_Syntax.improper_command @{command_spec "print_bnfs"} "print all BNFs"
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1284 |
(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
|
1285 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1286 |
val _ = |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1287 |
Outer_Syntax.local_theory_to_proof @{command_spec "bnf_def"} "define a BNF for an existing type"
|
|
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
|
1288 |
((parse_opt_binding_colon -- Parse.term -- |
| 49277 | 1289 |
(@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}) -- Parse.term --
|
| 49459 | 1290 |
(@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}) -- Scan.option Parse.term)
|
1291 |
>> bnf_def_cmd); |
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1292 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1293 |
end; |