author | wenzelm |
Sun, 30 Jan 2011 13:02:18 +0100 | |
changeset 41648 | 6d736d983d5c |
parent 41430 | 1aa23e9f2c87 |
child 42151 | 4da4fc77664b |
permissions | -rw-r--r-- |
37744 | 1 |
(* Title: HOLCF/Tools/Domain/domain_isomorphism.ML |
33774 | 2 |
Author: Brian Huffman |
3 |
||
4 |
Defines new types satisfying the given domain equations. |
|
5 |
*) |
|
6 |
||
7 |
signature DOMAIN_ISOMORPHISM = |
|
8 |
sig |
|
35489 | 9 |
val domain_isomorphism : |
35657
0537c34c6067
pass take_induct_info as an argument to comp_theorems
huffman
parents:
35654
diff
changeset
|
10 |
(string list * binding * mixfix * typ |
0537c34c6067
pass take_induct_info as an argument to comp_theorems
huffman
parents:
35654
diff
changeset
|
11 |
* (binding * binding) option) list -> |
0537c34c6067
pass take_induct_info as an argument to comp_theorems
huffman
parents:
35654
diff
changeset
|
12 |
theory -> |
0537c34c6067
pass take_induct_info as an argument to comp_theorems
huffman
parents:
35654
diff
changeset
|
13 |
(Domain_Take_Proofs.iso_info list |
0537c34c6067
pass take_induct_info as an argument to comp_theorems
huffman
parents:
35654
diff
changeset
|
14 |
* Domain_Take_Proofs.take_induct_info) * theory |
0537c34c6067
pass take_induct_info as an argument to comp_theorems
huffman
parents:
35654
diff
changeset
|
15 |
|
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
16 |
val define_map_functions : |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
17 |
(binding * Domain_Take_Proofs.iso_info) list -> |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
18 |
theory -> |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
19 |
{ |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
20 |
map_consts : term list, |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
21 |
map_apply_thms : thm list, |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
22 |
map_unfold_thms : thm list, |
41324
1383653efec3
make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents:
41292
diff
changeset
|
23 |
map_cont_thm : thm, |
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
24 |
deflation_map_thms : thm list |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
25 |
} |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
26 |
* theory |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
27 |
|
35489 | 28 |
val domain_isomorphism_cmd : |
34149
a0efb4754cb7
add 'morphisms' option to domain_isomorphism command
huffman
parents:
33971
diff
changeset
|
29 |
(string list * binding * mixfix * string * (binding * binding) option) list |
a0efb4754cb7
add 'morphisms' option to domain_isomorphism command
huffman
parents:
33971
diff
changeset
|
30 |
-> theory -> theory |
40218
f7d4d023a899
make domain package work with non-cpo argument types
huffman
parents:
40216
diff
changeset
|
31 |
|
40216
366309dfaf60
use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents:
40037
diff
changeset
|
32 |
val setup : theory -> theory |
40832 | 33 |
end |
33774 | 34 |
|
35514
a2cfa413eaab
move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
35498
diff
changeset
|
35 |
structure Domain_Isomorphism : DOMAIN_ISOMORPHISM = |
33774 | 36 |
struct |
37 |
||
40833
4f130bd9e17e
internal domain package proofs use cont2cont simproc instead of a fixed list of cont rules
huffman
parents:
40832
diff
changeset
|
38 |
val beta_ss = |
4f130bd9e17e
internal domain package proofs use cont2cont simproc instead of a fixed list of cont rules
huffman
parents:
40832
diff
changeset
|
39 |
HOL_basic_ss addsimps simp_thms addsimprocs [@{simproc beta_cfun_proc}] |
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
40 |
|
40832 | 41 |
fun is_cpo thy T = Sign.of_sort thy (T, @{sort cpo}) |
40218
f7d4d023a899
make domain package work with non-cpo argument types
huffman
parents:
40216
diff
changeset
|
42 |
|
33774 | 43 |
(******************************************************************************) |
33791 | 44 |
(******************************** theory data *********************************) |
45 |
(******************************************************************************) |
|
46 |
||
40216
366309dfaf60
use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents:
40037
diff
changeset
|
47 |
structure RepData = Named_Thms |
33791 | 48 |
( |
40216
366309dfaf60
use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents:
40037
diff
changeset
|
49 |
val name = "domain_defl_simps" |
366309dfaf60
use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents:
40037
diff
changeset
|
50 |
val description = "theorems like DEFL('a t) = t_defl$DEFL('a)" |
366309dfaf60
use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents:
40037
diff
changeset
|
51 |
) |
366309dfaf60
use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents:
40037
diff
changeset
|
52 |
|
366309dfaf60
use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents:
40037
diff
changeset
|
53 |
structure IsodeflData = Named_Thms |
35514
a2cfa413eaab
move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
35498
diff
changeset
|
54 |
( |
40216
366309dfaf60
use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents:
40037
diff
changeset
|
55 |
val name = "domain_isodefl" |
366309dfaf60
use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents:
40037
diff
changeset
|
56 |
val description = "theorems like isodefl d t ==> isodefl (foo_map$d) (foo_defl$t)" |
40832 | 57 |
) |
33809
033831fd9ef3
store map_ID thms in theory data; automate proofs of reach lemmas
huffman
parents:
33807
diff
changeset
|
58 |
|
40489 | 59 |
val setup = RepData.setup #> IsodeflData.setup |
33801 | 60 |
|
33791 | 61 |
|
62 |
(******************************************************************************) |
|
35475
979019ab92eb
move common functions into new file holcf_library.ML
huffman
parents:
35474
diff
changeset
|
63 |
(************************** building types and terms **************************) |
33774 | 64 |
(******************************************************************************) |
65 |
||
40832 | 66 |
open HOLCF_Library |
33774 | 67 |
|
40832 | 68 |
infixr 6 ->> |
69 |
infixr -->> |
|
33774 | 70 |
|
40832 | 71 |
val udomT = @{typ udom} |
41287
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
huffman
parents:
40833
diff
changeset
|
72 |
val deflT = @{typ "udom defl"} |
41292
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41287
diff
changeset
|
73 |
val udeflT = @{typ "udom u defl"} |
33774 | 74 |
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39974
diff
changeset
|
75 |
fun mk_DEFL T = |
40832 | 76 |
Const (@{const_name defl}, Term.itselfT T --> deflT) $ Logic.mk_type T |
33774 | 77 |
|
40487
1320a0747974
implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents:
40486
diff
changeset
|
78 |
fun dest_DEFL (Const (@{const_name defl}, _) $ t) = Logic.dest_type t |
40832 | 79 |
| dest_DEFL t = raise TERM ("dest_DEFL", [t]) |
40487
1320a0747974
implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents:
40486
diff
changeset
|
80 |
|
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
81 |
fun mk_LIFTDEFL T = |
41292
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41287
diff
changeset
|
82 |
Const (@{const_name liftdefl}, Term.itselfT T --> udeflT) $ Logic.mk_type T |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
83 |
|
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
84 |
fun dest_LIFTDEFL (Const (@{const_name liftdefl}, _) $ t) = Logic.dest_type t |
40832 | 85 |
| dest_LIFTDEFL t = raise TERM ("dest_LIFTDEFL", [t]) |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
86 |
|
40832 | 87 |
fun mk_u_defl t = mk_capply (@{const "u_defl"}, t) |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
88 |
|
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
89 |
fun mk_u_map t = |
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
90 |
let |
40832 | 91 |
val (T, U) = dest_cfunT (fastype_of t) |
92 |
val u_map_type = (T ->> U) ->> (mk_upT T ->> mk_upT U) |
|
93 |
val u_map_const = Const (@{const_name u_map}, u_map_type) |
|
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
94 |
in |
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
95 |
mk_capply (u_map_const, t) |
40832 | 96 |
end |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
97 |
|
40832 | 98 |
fun emb_const T = Const (@{const_name emb}, T ->> udomT) |
99 |
fun prj_const T = Const (@{const_name prj}, udomT ->> T) |
|
100 |
fun coerce_const (T, U) = mk_cfcomp (prj_const U, emb_const T) |
|
33778 | 101 |
|
33788 | 102 |
fun isodefl_const T = |
40832 | 103 |
Const (@{const_name isodefl}, (T ->> T) --> deflT --> HOLogic.boolT) |
33788 | 104 |
|
41292
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41287
diff
changeset
|
105 |
fun isodefl'_const T = |
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41287
diff
changeset
|
106 |
Const (@{const_name isodefl'}, (T ->> T) --> udeflT --> HOLogic.boolT) |
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41287
diff
changeset
|
107 |
|
35480
7a1f285cad25
domain_isomorphism package proves deflation rules for map functions
huffman
parents:
35479
diff
changeset
|
108 |
fun mk_deflation t = |
40832 | 109 |
Const (@{const_name deflation}, Term.fastype_of t --> boolT) $ t |
35480
7a1f285cad25
domain_isomorphism package proves deflation rules for map functions
huffman
parents:
35479
diff
changeset
|
110 |
|
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
111 |
(* splits a cterm into the right and lefthand sides of equality *) |
40832 | 112 |
fun dest_eqs t = HOLogic.dest_eq (HOLogic.dest_Trueprop t) |
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
113 |
|
40832 | 114 |
fun mk_eqs (t, u) = HOLogic.mk_Trueprop (HOLogic.mk_eq (t, u)) |
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
115 |
|
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
116 |
(******************************************************************************) |
35489 | 117 |
(****************************** isomorphism info ******************************) |
118 |
(******************************************************************************) |
|
119 |
||
35514
a2cfa413eaab
move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
35498
diff
changeset
|
120 |
fun deflation_abs_rep (info : Domain_Take_Proofs.iso_info) : thm = |
35489 | 121 |
let |
40832 | 122 |
val abs_iso = #abs_inverse info |
123 |
val rep_iso = #rep_inverse info |
|
124 |
val thm = @{thm deflation_abs_rep} OF [abs_iso, rep_iso] |
|
35489 | 125 |
in |
36241
2a4cec6bcae2
replace many uses of Drule.export_without_context with Drule.zero_var_indexes
huffman
parents:
35842
diff
changeset
|
126 |
Drule.zero_var_indexes thm |
35489 | 127 |
end |
128 |
||
129 |
(******************************************************************************) |
|
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
130 |
(*************** fixed-point definitions and unfolding theorems ***************) |
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
131 |
(******************************************************************************) |
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
132 |
|
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
133 |
fun mk_projs [] t = [] |
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
134 |
| mk_projs (x::[]) t = [(x, t)] |
40832 | 135 |
| mk_projs (x::xs) t = (x, mk_fst t) :: mk_projs xs (mk_snd t) |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
136 |
|
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
137 |
fun add_fixdefs |
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
138 |
(spec : (binding * term) list) |
41324
1383653efec3
make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents:
41292
diff
changeset
|
139 |
(thy : theory) : (thm list * thm list * thm) * theory = |
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
140 |
let |
40832 | 141 |
val binds = map fst spec |
142 |
val (lhss, rhss) = ListPair.unzip (map (dest_eqs o snd) spec) |
|
143 |
val functional = lambda_tuple lhss (mk_tuple rhss) |
|
144 |
val fixpoint = mk_fix (mk_cabs functional) |
|
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
145 |
|
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
146 |
(* project components of fixpoint *) |
40832 | 147 |
val projs = mk_projs lhss fixpoint |
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
148 |
|
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
149 |
(* convert parameters to lambda abstractions *) |
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
150 |
fun mk_eqn (lhs, rhs) = |
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
151 |
case lhs of |
40327 | 152 |
Const (@{const_name Rep_cfun}, _) $ f $ (x as Free _) => |
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
153 |
mk_eqn (f, big_lambda x rhs) |
40218
f7d4d023a899
make domain package work with non-cpo argument types
huffman
parents:
40216
diff
changeset
|
154 |
| f $ Const (@{const_name TYPE}, T) => |
f7d4d023a899
make domain package work with non-cpo argument types
huffman
parents:
40216
diff
changeset
|
155 |
mk_eqn (f, Abs ("t", T, rhs)) |
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
156 |
| Const _ => Logic.mk_equals (lhs, rhs) |
40832 | 157 |
| _ => raise TERM ("lhs not of correct form", [lhs, rhs]) |
158 |
val eqns = map mk_eqn projs |
|
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
159 |
|
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
160 |
(* register constant definitions *) |
33777 | 161 |
val (fixdef_thms, thy) = |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37744
diff
changeset
|
162 |
(Global_Theory.add_defs false o map Thm.no_attributes) |
40832 | 163 |
(map (Binding.suffix_name "_def") binds ~~ eqns) thy |
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
164 |
|
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
165 |
(* prove applied version of definitions *) |
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
166 |
fun prove_proj (lhs, rhs) = |
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
167 |
let |
40833
4f130bd9e17e
internal domain package proofs use cont2cont simproc instead of a fixed list of cont rules
huffman
parents:
40832
diff
changeset
|
168 |
val tac = rewrite_goals_tac fixdef_thms THEN (simp_tac beta_ss) 1 |
40832 | 169 |
val goal = Logic.mk_equals (lhs, rhs) |
170 |
in Goal.prove_global thy [] [] goal (K tac) end |
|
171 |
val proj_thms = map prove_proj projs |
|
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
172 |
|
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
173 |
(* mk_tuple lhss == fixpoint *) |
40832 | 174 |
fun pair_equalI (thm1, thm2) = @{thm Pair_equalI} OF [thm1, thm2] |
175 |
val tuple_fixdef_thm = foldr1 pair_equalI proj_thms |
|
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
176 |
|
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
177 |
val cont_thm = |
40833
4f130bd9e17e
internal domain package proofs use cont2cont simproc instead of a fixed list of cont rules
huffman
parents:
40832
diff
changeset
|
178 |
let |
4f130bd9e17e
internal domain package proofs use cont2cont simproc instead of a fixed list of cont rules
huffman
parents:
40832
diff
changeset
|
179 |
val prop = mk_trp (mk_cont functional) |
4f130bd9e17e
internal domain package proofs use cont2cont simproc instead of a fixed list of cont rules
huffman
parents:
40832
diff
changeset
|
180 |
val rules = Cont2ContData.get (ProofContext.init_global thy) |
4f130bd9e17e
internal domain package proofs use cont2cont simproc instead of a fixed list of cont rules
huffman
parents:
40832
diff
changeset
|
181 |
val tac = REPEAT_ALL_NEW (match_tac rules) 1 |
4f130bd9e17e
internal domain package proofs use cont2cont simproc instead of a fixed list of cont rules
huffman
parents:
40832
diff
changeset
|
182 |
in |
4f130bd9e17e
internal domain package proofs use cont2cont simproc instead of a fixed list of cont rules
huffman
parents:
40832
diff
changeset
|
183 |
Goal.prove_global thy [] [] prop (K tac) |
4f130bd9e17e
internal domain package proofs use cont2cont simproc instead of a fixed list of cont rules
huffman
parents:
40832
diff
changeset
|
184 |
end |
4f130bd9e17e
internal domain package proofs use cont2cont simproc instead of a fixed list of cont rules
huffman
parents:
40832
diff
changeset
|
185 |
|
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
186 |
val tuple_unfold_thm = |
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
187 |
(@{thm def_cont_fix_eq} OF [tuple_fixdef_thm, cont_thm]) |
40832 | 188 |
|> Local_Defs.unfold (ProofContext.init_global thy) @{thms split_conv} |
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
189 |
|
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
190 |
fun mk_unfold_thms [] thm = [] |
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
191 |
| mk_unfold_thms (n::[]) thm = [(n, thm)] |
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
192 |
| mk_unfold_thms (n::ns) thm = let |
40832 | 193 |
val thmL = thm RS @{thm Pair_eqD1} |
194 |
val thmR = thm RS @{thm Pair_eqD2} |
|
195 |
in (n, thmL) :: mk_unfold_thms ns thmR end |
|
196 |
val unfold_binds = map (Binding.suffix_name "_unfold") binds |
|
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
197 |
|
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
198 |
(* register unfold theorems *) |
33777 | 199 |
val (unfold_thms, thy) = |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37744
diff
changeset
|
200 |
(Global_Theory.add_thms o map (Thm.no_attributes o apsnd Drule.zero_var_indexes)) |
40832 | 201 |
(mk_unfold_thms unfold_binds tuple_unfold_thm) thy |
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
202 |
in |
41324
1383653efec3
make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents:
41292
diff
changeset
|
203 |
((proj_thms, unfold_thms, cont_thm), thy) |
40832 | 204 |
end |
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
205 |
|
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
206 |
|
33774 | 207 |
(******************************************************************************) |
33790 | 208 |
(****************** deflation combinators and map functions *******************) |
209 |
(******************************************************************************) |
|
33774 | 210 |
|
211 |
fun defl_of_typ |
|
40218
f7d4d023a899
make domain package work with non-cpo argument types
huffman
parents:
40216
diff
changeset
|
212 |
(thy : theory) |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
213 |
(tab1 : (typ * term) list) |
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
214 |
(tab2 : (typ * term) list) |
33774 | 215 |
(T : typ) : term = |
216 |
let |
|
40832 | 217 |
val defl_simps = RepData.get (ProofContext.init_global thy) |
218 |
val rules = map (Thm.concl_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq) defl_simps |
|
219 |
val rules' = map (apfst mk_DEFL) tab1 @ map (apfst mk_LIFTDEFL) tab2 |
|
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
220 |
fun proc1 t = |
40487
1320a0747974
implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents:
40486
diff
changeset
|
221 |
(case dest_DEFL t of |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
222 |
TFree (a, _) => SOME (Free ("d" ^ Library.unprefix "'" a, deflT)) |
40832 | 223 |
| _ => NONE) handle TERM _ => NONE |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
224 |
fun proc2 t = |
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
225 |
(case dest_LIFTDEFL t of |
41292
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41287
diff
changeset
|
226 |
TFree (a, _) => SOME (Free ("p" ^ Library.unprefix "'" a, udeflT)) |
40832 | 227 |
| _ => NONE) handle TERM _ => NONE |
40487
1320a0747974
implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents:
40486
diff
changeset
|
228 |
in |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
229 |
Pattern.rewrite_term thy (rules @ rules') [proc1, proc2] (mk_DEFL T) |
40832 | 230 |
end |
33774 | 231 |
|
232 |
(******************************************************************************) |
|
35477 | 233 |
(********************* declaring definitions and theorems *********************) |
234 |
(******************************************************************************) |
|
235 |
||
236 |
fun define_const |
|
237 |
(bind : binding, rhs : term) |
|
238 |
(thy : theory) |
|
239 |
: (term * thm) * theory = |
|
240 |
let |
|
40832 | 241 |
val typ = Term.fastype_of rhs |
242 |
val (const, thy) = Sign.declare_const ((bind, typ), NoSyn) thy |
|
243 |
val eqn = Logic.mk_equals (const, rhs) |
|
244 |
val def = Thm.no_attributes (Binding.suffix_name "_def" bind, eqn) |
|
245 |
val (def_thm, thy) = yield_singleton (Global_Theory.add_defs false) def thy |
|
35477 | 246 |
in |
247 |
((const, def_thm), thy) |
|
40832 | 248 |
end |
35477 | 249 |
|
35773 | 250 |
fun add_qualified_thm name (dbind, thm) = |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37744
diff
changeset
|
251 |
yield_singleton Global_Theory.add_thms |
40832 | 252 |
((Binding.qualified true name dbind, thm), []) |
35489 | 253 |
|
254 |
(******************************************************************************) |
|
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
255 |
(*************************** defining map functions ***************************) |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
256 |
(******************************************************************************) |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
257 |
|
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
258 |
fun define_map_functions |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
259 |
(spec : (binding * Domain_Take_Proofs.iso_info) list) |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
260 |
(thy : theory) = |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
261 |
let |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
262 |
|
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
263 |
(* retrieve components of spec *) |
40832 | 264 |
val dbinds = map fst spec |
265 |
val iso_infos = map snd spec |
|
266 |
val dom_eqns = map (fn x => (#absT x, #repT x)) iso_infos |
|
267 |
val rep_abs_consts = map (fn x => (#rep_const x, #abs_const x)) iso_infos |
|
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
268 |
|
40218
f7d4d023a899
make domain package work with non-cpo argument types
huffman
parents:
40216
diff
changeset
|
269 |
fun mapT (T as Type (_, Ts)) = |
f7d4d023a899
make domain package work with non-cpo argument types
huffman
parents:
40216
diff
changeset
|
270 |
(map (fn T => T ->> T) (filter (is_cpo thy) Ts)) -->> (T ->> T) |
40832 | 271 |
| mapT T = T ->> T |
40218
f7d4d023a899
make domain package work with non-cpo argument types
huffman
parents:
40216
diff
changeset
|
272 |
|
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
273 |
(* declare map functions *) |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
274 |
fun declare_map_const (tbind, (lhsT, rhsT)) thy = |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
275 |
let |
40832 | 276 |
val map_type = mapT lhsT |
277 |
val map_bind = Binding.suffix_name "_map" tbind |
|
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
278 |
in |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
279 |
Sign.declare_const ((map_bind, map_type), NoSyn) thy |
40832 | 280 |
end |
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
281 |
val (map_consts, thy) = thy |> |
40832 | 282 |
fold_map declare_map_const (dbinds ~~ dom_eqns) |
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
283 |
|
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
284 |
(* defining equations for map functions *) |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
285 |
local |
40832 | 286 |
fun unprime a = Library.unprefix "'" a |
287 |
fun mapvar T = Free (unprime (fst (dest_TFree T)), T ->> T) |
|
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
288 |
fun map_lhs (map_const, lhsT) = |
40832 | 289 |
(lhsT, list_ccomb (map_const, map mapvar (filter (is_cpo thy) (snd (dest_Type lhsT))))) |
290 |
val tab1 = map map_lhs (map_consts ~~ map fst dom_eqns) |
|
291 |
val Ts = (snd o dest_Type o fst o hd) dom_eqns |
|
292 |
val tab = (Ts ~~ map mapvar Ts) @ tab1 |
|
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
293 |
fun mk_map_spec (((rep_const, abs_const), map_const), (lhsT, rhsT)) = |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
294 |
let |
40832 | 295 |
val lhs = Domain_Take_Proofs.map_of_typ thy tab lhsT |
296 |
val body = Domain_Take_Proofs.map_of_typ thy tab rhsT |
|
297 |
val rhs = mk_cfcomp (abs_const, mk_cfcomp (body, rep_const)) |
|
298 |
in mk_eqs (lhs, rhs) end |
|
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
299 |
in |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
300 |
val map_specs = |
40832 | 301 |
map mk_map_spec (rep_abs_consts ~~ map_consts ~~ dom_eqns) |
302 |
end |
|
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
303 |
|
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
304 |
(* register recursive definition of map functions *) |
40832 | 305 |
val map_binds = map (Binding.suffix_name "_map") dbinds |
41324
1383653efec3
make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents:
41292
diff
changeset
|
306 |
val ((map_apply_thms, map_unfold_thms, map_cont_thm), thy) = |
40832 | 307 |
add_fixdefs (map_binds ~~ map_specs) thy |
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
308 |
|
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
309 |
(* prove deflation theorems for map functions *) |
40832 | 310 |
val deflation_abs_rep_thms = map deflation_abs_rep iso_infos |
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
311 |
val deflation_map_thm = |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
312 |
let |
40832 | 313 |
fun unprime a = Library.unprefix "'" a |
314 |
fun mk_f T = Free (unprime (fst (dest_TFree T)), T ->> T) |
|
315 |
fun mk_assm T = mk_trp (mk_deflation (mk_f T)) |
|
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
316 |
fun mk_goal (map_const, (lhsT, rhsT)) = |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
317 |
let |
40832 | 318 |
val (_, Ts) = dest_Type lhsT |
319 |
val map_term = list_ccomb (map_const, map mk_f (filter (is_cpo thy) Ts)) |
|
320 |
in mk_deflation map_term end |
|
321 |
val assms = (map mk_assm o filter (is_cpo thy) o snd o dest_Type o fst o hd) dom_eqns |
|
322 |
val goals = map mk_goal (map_consts ~~ dom_eqns) |
|
323 |
val goal = mk_trp (foldr1 HOLogic.mk_conj goals) |
|
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
324 |
val adm_rules = |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
325 |
@{thms adm_conj adm_subst [OF _ adm_deflation] |
40832 | 326 |
cont2cont_fst cont2cont_snd cont_id} |
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
327 |
val bottom_rules = |
41430
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
huffman
parents:
41324
diff
changeset
|
328 |
@{thms fst_strict snd_strict deflation_bottom simp_thms} |
41324
1383653efec3
make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents:
41292
diff
changeset
|
329 |
val tuple_rules = |
1383653efec3
make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents:
41292
diff
changeset
|
330 |
@{thms split_def fst_conv snd_conv} |
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
331 |
val deflation_rules = |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
332 |
@{thms conjI deflation_ID} |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
333 |
@ deflation_abs_rep_thms |
40832 | 334 |
@ Domain_Take_Proofs.get_deflation_thms thy |
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
335 |
in |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
336 |
Goal.prove_global thy [] assms goal (fn {prems, ...} => |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
337 |
EVERY |
41324
1383653efec3
make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents:
41292
diff
changeset
|
338 |
[rewrite_goals_tac map_apply_thms, |
1383653efec3
make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents:
41292
diff
changeset
|
339 |
rtac (map_cont_thm RS @{thm cont_fix_ind}) 1, |
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
340 |
REPEAT (resolve_tac adm_rules 1), |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
341 |
simp_tac (HOL_basic_ss addsimps bottom_rules) 1, |
41324
1383653efec3
make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents:
41292
diff
changeset
|
342 |
simp_tac (HOL_basic_ss addsimps tuple_rules) 1, |
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
343 |
REPEAT (etac @{thm conjE} 1), |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
344 |
REPEAT (resolve_tac (deflation_rules @ prems) 1 ORELSE atac 1)]) |
40832 | 345 |
end |
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
346 |
fun conjuncts [] thm = [] |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
347 |
| conjuncts (n::[]) thm = [(n, thm)] |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
348 |
| conjuncts (n::ns) thm = let |
40832 | 349 |
val thmL = thm RS @{thm conjunct1} |
350 |
val thmR = thm RS @{thm conjunct2} |
|
351 |
in (n, thmL):: conjuncts ns thmR end |
|
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
352 |
val deflation_map_binds = dbinds |> |
40832 | 353 |
map (Binding.prefix_name "deflation_" o Binding.suffix_name "_map") |
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
354 |
val (deflation_map_thms, thy) = thy |> |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37744
diff
changeset
|
355 |
(Global_Theory.add_thms o map (Thm.no_attributes o apsnd Drule.zero_var_indexes)) |
40832 | 356 |
(conjuncts deflation_map_binds deflation_map_thm) |
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
357 |
|
40737
2037021f034f
remove map function names from domain package theory data
huffman
parents:
40575
diff
changeset
|
358 |
(* register indirect recursion in theory data *) |
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
359 |
local |
40737
2037021f034f
remove map function names from domain package theory data
huffman
parents:
40575
diff
changeset
|
360 |
fun register_map (dname, args) = |
40832 | 361 |
Domain_Take_Proofs.add_rec_type (dname, args) |
362 |
val dnames = map (fst o dest_Type o fst) dom_eqns |
|
363 |
val map_names = map (fst o dest_Const) map_consts |
|
364 |
fun args (T, _) = case T of Type (_, Ts) => map (is_cpo thy) Ts | _ => [] |
|
365 |
val argss = map args dom_eqns |
|
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
366 |
in |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
367 |
val thy = |
40832 | 368 |
fold register_map (dnames ~~ argss) thy |
369 |
end |
|
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
370 |
|
40216
366309dfaf60
use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents:
40037
diff
changeset
|
371 |
(* register deflation theorems *) |
40832 | 372 |
val thy = fold Domain_Take_Proofs.add_deflation_thm deflation_map_thms thy |
40216
366309dfaf60
use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents:
40037
diff
changeset
|
373 |
|
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
374 |
val result = |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
375 |
{ |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
376 |
map_consts = map_consts, |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
377 |
map_apply_thms = map_apply_thms, |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
378 |
map_unfold_thms = map_unfold_thms, |
41324
1383653efec3
make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents:
41292
diff
changeset
|
379 |
map_cont_thm = map_cont_thm, |
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
380 |
deflation_map_thms = deflation_map_thms |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
381 |
} |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
382 |
in |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
383 |
(result, thy) |
40832 | 384 |
end |
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
385 |
|
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
386 |
(******************************************************************************) |
35477 | 387 |
(******************************* main function ********************************) |
388 |
(******************************************************************************) |
|
33774 | 389 |
|
390 |
fun read_typ thy str sorts = |
|
391 |
let |
|
36610
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents:
36241
diff
changeset
|
392 |
val ctxt = ProofContext.init_global thy |
40832 | 393 |
|> fold (Variable.declare_typ o TFree) sorts |
394 |
val T = Syntax.read_typ ctxt str |
|
395 |
in (T, Term.add_tfreesT T sorts) end |
|
33774 | 396 |
|
397 |
fun cert_typ sign raw_T sorts = |
|
398 |
let |
|
399 |
val T = Type.no_tvars (Sign.certify_typ sign raw_T) |
|
40832 | 400 |
handle TYPE (msg, _, _) => error msg |
401 |
val sorts' = Term.add_tfreesT T sorts |
|
33774 | 402 |
val _ = |
403 |
case duplicates (op =) (map fst sorts') of |
|
404 |
[] => () |
|
405 |
| dups => error ("Inconsistent sort constraints for " ^ commas dups) |
|
40832 | 406 |
in (T, sorts') end |
33774 | 407 |
|
408 |
fun gen_domain_isomorphism |
|
409 |
(prep_typ: theory -> 'a -> (string * sort) list -> typ * (string * sort) list) |
|
34149
a0efb4754cb7
add 'morphisms' option to domain_isomorphism command
huffman
parents:
33971
diff
changeset
|
410 |
(doms_raw: (string list * binding * mixfix * 'a * (binding * binding) option) list) |
33774 | 411 |
(thy: theory) |
35657
0537c34c6067
pass take_induct_info as an argument to comp_theorems
huffman
parents:
35654
diff
changeset
|
412 |
: (Domain_Take_Proofs.iso_info list |
0537c34c6067
pass take_induct_info as an argument to comp_theorems
huffman
parents:
35654
diff
changeset
|
413 |
* Domain_Take_Proofs.take_induct_info) * theory = |
33774 | 414 |
let |
40832 | 415 |
val _ = Theory.requires thy "Domain" "domain isomorphisms" |
33774 | 416 |
|
417 |
(* this theory is used just for parsing *) |
|
418 |
val tmp_thy = thy |> |
|
419 |
Theory.copy |> |
|
40486 | 420 |
Sign.add_types (map (fn (tvs, tbind, mx, _, morphs) => |
40832 | 421 |
(tbind, length tvs, mx)) doms_raw) |
33774 | 422 |
|
34149
a0efb4754cb7
add 'morphisms' option to domain_isomorphism command
huffman
parents:
33971
diff
changeset
|
423 |
fun prep_dom thy (vs, t, mx, typ_raw, morphs) sorts = |
33774 | 424 |
let val (typ, sorts') = prep_typ thy typ_raw sorts |
40832 | 425 |
in ((vs, t, mx, typ, morphs), sorts') end |
33774 | 426 |
|
34149
a0efb4754cb7
add 'morphisms' option to domain_isomorphism command
huffman
parents:
33971
diff
changeset
|
427 |
val (doms : (string list * binding * mixfix * typ * (binding * binding) option) list, |
33774 | 428 |
sorts : (string * sort) list) = |
40832 | 429 |
fold_map (prep_dom tmp_thy) doms_raw [] |
33774 | 430 |
|
40486 | 431 |
(* lookup function for sorts of type variables *) |
40832 | 432 |
fun the_sort v = the (AList.lookup (op =) sorts v) |
40486 | 433 |
|
40490
05be0c37db1d
add bifiniteness check for domain_isomorphism command
huffman
parents:
40489
diff
changeset
|
434 |
(* declare arities in temporary theory *) |
40487
1320a0747974
implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents:
40486
diff
changeset
|
435 |
val tmp_thy = |
1320a0747974
implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents:
40486
diff
changeset
|
436 |
let |
1320a0747974
implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents:
40486
diff
changeset
|
437 |
fun arity (vs, tbind, mx, _, _) = |
40832 | 438 |
(Sign.full_name thy tbind, map the_sort vs, @{sort "domain"}) |
40487
1320a0747974
implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents:
40486
diff
changeset
|
439 |
in |
1320a0747974
implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents:
40486
diff
changeset
|
440 |
fold AxClass.axiomatize_arity (map arity doms) tmp_thy |
40832 | 441 |
end |
40487
1320a0747974
implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents:
40486
diff
changeset
|
442 |
|
40490
05be0c37db1d
add bifiniteness check for domain_isomorphism command
huffman
parents:
40489
diff
changeset
|
443 |
(* check bifiniteness of right-hand sides *) |
05be0c37db1d
add bifiniteness check for domain_isomorphism command
huffman
parents:
40489
diff
changeset
|
444 |
fun check_rhs (vs, tbind, mx, rhs, morphs) = |
40497 | 445 |
if Sign.of_sort tmp_thy (rhs, @{sort "domain"}) then () |
446 |
else error ("Type not of sort domain: " ^ |
|
40832 | 447 |
quote (Syntax.string_of_typ_global tmp_thy rhs)) |
448 |
val _ = map check_rhs doms |
|
40490
05be0c37db1d
add bifiniteness check for domain_isomorphism command
huffman
parents:
40489
diff
changeset
|
449 |
|
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
450 |
(* domain equations *) |
34149
a0efb4754cb7
add 'morphisms' option to domain_isomorphism command
huffman
parents:
33971
diff
changeset
|
451 |
fun mk_dom_eqn (vs, tbind, mx, rhs, morphs) = |
40832 | 452 |
let fun arg v = TFree (v, the_sort v) |
453 |
in (Type (Sign.full_name tmp_thy tbind, map arg vs), rhs) end |
|
454 |
val dom_eqns = map mk_dom_eqn doms |
|
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
455 |
|
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
456 |
(* check for valid type parameters *) |
40832 | 457 |
val (tyvars, _, _, _, _) = hd doms |
34149
a0efb4754cb7
add 'morphisms' option to domain_isomorphism command
huffman
parents:
33971
diff
changeset
|
458 |
val new_doms = map (fn (tvs, tname, mx, _, _) => |
33774 | 459 |
let val full_tname = Sign.full_name tmp_thy tname |
460 |
in |
|
461 |
(case duplicates (op =) tvs of |
|
462 |
[] => |
|
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
463 |
if eq_set (op =) (tyvars, tvs) then (full_tname, tvs) |
33774 | 464 |
else error ("Mutually recursive domains must have same type parameters") |
465 |
| dups => error ("Duplicate parameter(s) for domain " ^ quote (Binding.str_of tname) ^ |
|
466 |
" : " ^ commas dups)) |
|
40832 | 467 |
end) doms |
468 |
val dbinds = map (fn (_, dbind, _, _, _) => dbind) doms |
|
469 |
val morphs = map (fn (_, _, _, _, morphs) => morphs) doms |
|
33774 | 470 |
|
40218
f7d4d023a899
make domain package work with non-cpo argument types
huffman
parents:
40216
diff
changeset
|
471 |
(* determine deflation combinator arguments *) |
40832 | 472 |
val lhsTs : typ list = map fst dom_eqns |
473 |
val defl_rec = Free ("t", mk_tupleT (map (K deflT) lhsTs)) |
|
474 |
val defl_recs = mk_projs lhsTs defl_rec |
|
475 |
val defl_recs' = map (apsnd mk_u_defl) defl_recs |
|
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
476 |
fun defl_body (_, _, _, rhsT, _) = |
40832 | 477 |
defl_of_typ tmp_thy defl_recs defl_recs' rhsT |
478 |
val functional = Term.lambda defl_rec (mk_tuple (map defl_body doms)) |
|
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
479 |
|
40832 | 480 |
val tfrees = map fst (Term.add_tfrees functional []) |
481 |
val frees = map fst (Term.add_frees functional []) |
|
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
482 |
fun get_defl_flags (vs, _, _, _, _) = |
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
483 |
let |
40832 | 484 |
fun argT v = TFree (v, the_sort v) |
485 |
fun mk_d v = "d" ^ Library.unprefix "'" v |
|
486 |
fun mk_p v = "p" ^ Library.unprefix "'" v |
|
487 |
val args = maps (fn v => [(mk_d v, mk_DEFL (argT v)), (mk_p v, mk_LIFTDEFL (argT v))]) vs |
|
488 |
val typeTs = map argT (filter (member (op =) tfrees) vs) |
|
489 |
val defl_args = map snd (filter (member (op =) frees o fst) args) |
|
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
490 |
in |
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
491 |
(typeTs, defl_args) |
40832 | 492 |
end |
493 |
val defl_flagss = map get_defl_flags doms |
|
40218
f7d4d023a899
make domain package work with non-cpo argument types
huffman
parents:
40216
diff
changeset
|
494 |
|
33784
7e434813752f
change naming convention for deflation combinators
huffman
parents:
33782
diff
changeset
|
495 |
(* declare deflation combinator constants *) |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
496 |
fun declare_defl_const ((typeTs, defl_args), (_, tbind, _, _, _)) thy = |
33774 | 497 |
let |
40832 | 498 |
val defl_bind = Binding.suffix_name "_defl" tbind |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
499 |
val defl_type = |
41292
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41287
diff
changeset
|
500 |
map Term.itselfT typeTs ---> map fastype_of defl_args -->> deflT |
33774 | 501 |
in |
33784
7e434813752f
change naming convention for deflation combinators
huffman
parents:
33782
diff
changeset
|
502 |
Sign.declare_const ((defl_bind, defl_type), NoSyn) thy |
40832 | 503 |
end |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
504 |
val (defl_consts, thy) = |
40832 | 505 |
fold_map declare_defl_const (defl_flagss ~~ doms) thy |
33774 | 506 |
|
33775
7a1518c42c56
cleaned up; factored out fixed-point definition code
huffman
parents:
33774
diff
changeset
|
507 |
(* defining equations for type combinators *) |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
508 |
fun mk_defl_term (defl_const, (typeTs, defl_args)) = |
40487
1320a0747974
implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents:
40486
diff
changeset
|
509 |
let |
40832 | 510 |
val type_args = map Logic.mk_type typeTs |
40487
1320a0747974
implement defl_of_typ using Pattern.rewrite_term instead of DeflData theory data
huffman
parents:
40486
diff
changeset
|
511 |
in |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
512 |
list_ccomb (list_comb (defl_const, type_args), defl_args) |
40832 | 513 |
end |
514 |
val defl_terms = map mk_defl_term (defl_consts ~~ defl_flagss) |
|
515 |
val defl_tab = map fst dom_eqns ~~ defl_terms |
|
516 |
val defl_tab' = map fst dom_eqns ~~ map mk_u_defl defl_terms |
|
33776 | 517 |
fun mk_defl_spec (lhsT, rhsT) = |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
518 |
mk_eqs (defl_of_typ tmp_thy defl_tab defl_tab' lhsT, |
40832 | 519 |
defl_of_typ tmp_thy defl_tab defl_tab' rhsT) |
520 |
val defl_specs = map mk_defl_spec dom_eqns |
|
33774 | 521 |
|
33784
7e434813752f
change naming convention for deflation combinators
huffman
parents:
33782
diff
changeset
|
522 |
(* register recursive definition of deflation combinators *) |
40832 | 523 |
val defl_binds = map (Binding.suffix_name "_defl") dbinds |
41324
1383653efec3
make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents:
41292
diff
changeset
|
524 |
val ((defl_apply_thms, defl_unfold_thms, defl_cont_thm), thy) = |
40832 | 525 |
add_fixdefs (defl_binds ~~ defl_specs) thy |
33774 | 526 |
|
33776 | 527 |
(* define types using deflation combinators *) |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
528 |
fun make_repdef ((vs, tbind, mx, _, _), defl) thy = |
33774 | 529 |
let |
40832 | 530 |
val spec = (tbind, map (rpair dummyS) vs, mx) |
40494
db8a09daba7b
add class liftdomain, for bifinite domains where DEFL('a u) = u_defl('a)
huffman
parents:
40492
diff
changeset
|
531 |
val ((_, _, _, {DEFL, liftemb_def, liftprj_def, ...}), thy) = |
40832 | 532 |
Domaindef.add_domaindef false NONE spec defl NONE thy |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
533 |
(* declare domain_defl_simps rules *) |
40832 | 534 |
val thy = Context.theory_map (RepData.add_thm DEFL) thy |
33774 | 535 |
in |
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39974
diff
changeset
|
536 |
(DEFL, thy) |
40832 | 537 |
end |
538 |
val (DEFL_thms, thy) = fold_map make_repdef (doms ~~ defl_terms) thy |
|
33776 | 539 |
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39974
diff
changeset
|
540 |
(* prove DEFL equations *) |
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39974
diff
changeset
|
541 |
fun mk_DEFL_eq_thm (lhsT, rhsT) = |
33776 | 542 |
let |
40832 | 543 |
val goal = mk_eqs (mk_DEFL lhsT, mk_DEFL rhsT) |
544 |
val DEFL_simps = RepData.get (ProofContext.init_global thy) |
|
33776 | 545 |
val tac = |
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39974
diff
changeset
|
546 |
rewrite_goals_tac (map mk_meta_eq DEFL_simps) |
40832 | 547 |
THEN TRY (resolve_tac defl_unfold_thms 1) |
33776 | 548 |
in |
33777 | 549 |
Goal.prove_global thy [] [] goal (K tac) |
40832 | 550 |
end |
551 |
val DEFL_eq_thms = map mk_DEFL_eq_thm dom_eqns |
|
33782 | 552 |
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39974
diff
changeset
|
553 |
(* register DEFL equations *) |
40832 | 554 |
val DEFL_eq_binds = map (Binding.prefix_name "DEFL_eq_") dbinds |
33782 | 555 |
val (_, thy) = thy |> |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37744
diff
changeset
|
556 |
(Global_Theory.add_thms o map Thm.no_attributes) |
40832 | 557 |
(DEFL_eq_binds ~~ DEFL_eq_thms) |
33776 | 558 |
|
33778 | 559 |
(* define rep/abs functions *) |
34149
a0efb4754cb7
add 'morphisms' option to domain_isomorphism command
huffman
parents:
33971
diff
changeset
|
560 |
fun mk_rep_abs ((tbind, morphs), (lhsT, rhsT)) thy = |
33778 | 561 |
let |
40832 | 562 |
val rep_bind = Binding.suffix_name "_rep" tbind |
563 |
val abs_bind = Binding.suffix_name "_abs" tbind |
|
35477 | 564 |
val ((rep_const, rep_def), thy) = |
40832 | 565 |
define_const (rep_bind, coerce_const (lhsT, rhsT)) thy |
35477 | 566 |
val ((abs_const, abs_def), thy) = |
40832 | 567 |
define_const (abs_bind, coerce_const (rhsT, lhsT)) thy |
33778 | 568 |
in |
33785
2f2d9eb37084
automate definition of map functions; remove unused code
huffman
parents:
33784
diff
changeset
|
569 |
(((rep_const, abs_const), (rep_def, abs_def)), thy) |
40832 | 570 |
end |
33785
2f2d9eb37084
automate definition of map functions; remove unused code
huffman
parents:
33784
diff
changeset
|
571 |
val ((rep_abs_consts, rep_abs_defs), thy) = thy |
35773 | 572 |
|> fold_map mk_rep_abs (dbinds ~~ morphs ~~ dom_eqns) |
40832 | 573 |
|>> ListPair.unzip |
33782 | 574 |
|
575 |
(* prove isomorphism and isodefl rules *) |
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39974
diff
changeset
|
576 |
fun mk_iso_thms ((tbind, DEFL_eq), (rep_def, abs_def)) thy = |
33782 | 577 |
let |
35773 | 578 |
fun make thm = |
40832 | 579 |
Drule.zero_var_indexes (thm OF [DEFL_eq, abs_def, rep_def]) |
580 |
val rep_iso_thm = make @{thm domain_rep_iso} |
|
581 |
val abs_iso_thm = make @{thm domain_abs_iso} |
|
582 |
val isodefl_thm = make @{thm isodefl_abs_rep} |
|
35773 | 583 |
val thy = thy |
584 |
|> snd o add_qualified_thm "rep_iso" (tbind, rep_iso_thm) |
|
585 |
|> snd o add_qualified_thm "abs_iso" (tbind, abs_iso_thm) |
|
40832 | 586 |
|> snd o add_qualified_thm "isodefl_abs_rep" (tbind, isodefl_thm) |
33782 | 587 |
in |
588 |
(((rep_iso_thm, abs_iso_thm), isodefl_thm), thy) |
|
40832 | 589 |
end |
35489 | 590 |
val ((iso_thms, isodefl_abs_rep_thms), thy) = |
35480
7a1f285cad25
domain_isomorphism package proves deflation rules for map functions
huffman
parents:
35479
diff
changeset
|
591 |
thy |
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39974
diff
changeset
|
592 |
|> fold_map mk_iso_thms (dbinds ~~ DEFL_eq_thms ~~ rep_abs_defs) |
40832 | 593 |
|>> ListPair.unzip |
33778 | 594 |
|
35467
561d8e98d9d3
domain_isomorphism function returns iso_info record
huffman
parents:
35021
diff
changeset
|
595 |
(* collect info about rep/abs *) |
35514
a2cfa413eaab
move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
35498
diff
changeset
|
596 |
val iso_infos : Domain_Take_Proofs.iso_info list = |
35467
561d8e98d9d3
domain_isomorphism function returns iso_info record
huffman
parents:
35021
diff
changeset
|
597 |
let |
561d8e98d9d3
domain_isomorphism function returns iso_info record
huffman
parents:
35021
diff
changeset
|
598 |
fun mk_info (((lhsT, rhsT), (repC, absC)), (rep_iso, abs_iso)) = |
561d8e98d9d3
domain_isomorphism function returns iso_info record
huffman
parents:
35021
diff
changeset
|
599 |
{ |
561d8e98d9d3
domain_isomorphism function returns iso_info record
huffman
parents:
35021
diff
changeset
|
600 |
repT = rhsT, |
561d8e98d9d3
domain_isomorphism function returns iso_info record
huffman
parents:
35021
diff
changeset
|
601 |
absT = lhsT, |
561d8e98d9d3
domain_isomorphism function returns iso_info record
huffman
parents:
35021
diff
changeset
|
602 |
rep_const = repC, |
561d8e98d9d3
domain_isomorphism function returns iso_info record
huffman
parents:
35021
diff
changeset
|
603 |
abs_const = absC, |
561d8e98d9d3
domain_isomorphism function returns iso_info record
huffman
parents:
35021
diff
changeset
|
604 |
rep_inverse = rep_iso, |
561d8e98d9d3
domain_isomorphism function returns iso_info record
huffman
parents:
35021
diff
changeset
|
605 |
abs_inverse = abs_iso |
40832 | 606 |
} |
35467
561d8e98d9d3
domain_isomorphism function returns iso_info record
huffman
parents:
35021
diff
changeset
|
607 |
in |
561d8e98d9d3
domain_isomorphism function returns iso_info record
huffman
parents:
35021
diff
changeset
|
608 |
map mk_info (dom_eqns ~~ rep_abs_consts ~~ iso_thms) |
561d8e98d9d3
domain_isomorphism function returns iso_info record
huffman
parents:
35021
diff
changeset
|
609 |
end |
561d8e98d9d3
domain_isomorphism function returns iso_info record
huffman
parents:
35021
diff
changeset
|
610 |
|
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
611 |
(* definitions and proofs related to map functions *) |
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
612 |
val (map_info, thy) = |
40832 | 613 |
define_map_functions (dbinds ~~ iso_infos) thy |
35791
dc175fe29326
separate map-related code into new function define_map_functions
huffman
parents:
35773
diff
changeset
|
614 |
val { map_consts, map_apply_thms, map_unfold_thms, |
41324
1383653efec3
make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents:
41292
diff
changeset
|
615 |
map_cont_thm, deflation_map_thms } = map_info |
33788 | 616 |
|
617 |
(* prove isodefl rules for map functions *) |
|
618 |
val isodefl_thm = |
|
619 |
let |
|
40832 | 620 |
fun unprime a = Library.unprefix "'" a |
621 |
fun mk_d T = Free ("d" ^ unprime (fst (dest_TFree T)), deflT) |
|
41292
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41287
diff
changeset
|
622 |
fun mk_p T = Free ("p" ^ unprime (fst (dest_TFree T)), udeflT) |
40832 | 623 |
fun mk_f T = Free ("f" ^ unprime (fst (dest_TFree T)), T ->> T) |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
624 |
fun mk_assm t = |
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
625 |
case try dest_LIFTDEFL t of |
41292
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41287
diff
changeset
|
626 |
SOME T => mk_trp (isodefl'_const T $ mk_f T $ mk_p T) |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
627 |
| NONE => |
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
628 |
let val T = dest_DEFL t |
40832 | 629 |
in mk_trp (isodefl_const T $ mk_f T $ mk_d T) end |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
630 |
fun mk_goal (map_const, (T, rhsT)) = |
33788 | 631 |
let |
40832 | 632 |
val (_, Ts) = dest_Type T |
633 |
val map_term = list_ccomb (map_const, map mk_f (filter (is_cpo thy) Ts)) |
|
634 |
val defl_term = defl_of_typ thy (Ts ~~ map mk_d Ts) (Ts ~~ map mk_p Ts) T |
|
635 |
in isodefl_const T $ map_term $ defl_term end |
|
636 |
val assms = (map mk_assm o snd o hd) defl_flagss |
|
637 |
val goals = map mk_goal (map_consts ~~ dom_eqns) |
|
638 |
val goal = mk_trp (foldr1 HOLogic.mk_conj goals) |
|
33788 | 639 |
val adm_rules = |
40832 | 640 |
@{thms adm_conj adm_isodefl cont2cont_fst cont2cont_snd cont_id} |
33788 | 641 |
val bottom_rules = |
40832 | 642 |
@{thms fst_strict snd_strict isodefl_bottom simp_thms} |
41324
1383653efec3
make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents:
41292
diff
changeset
|
643 |
val tuple_rules = |
1383653efec3
make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents:
41292
diff
changeset
|
644 |
@{thms split_def fst_conv snd_conv} |
40832 | 645 |
val map_ID_thms = Domain_Take_Proofs.get_map_ID_thms thy |
646 |
val map_ID_simps = map (fn th => th RS sym) map_ID_thms |
|
33788 | 647 |
val isodefl_rules = |
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40490
diff
changeset
|
648 |
@{thms conjI isodefl_ID_DEFL isodefl_LIFTDEFL} |
33791 | 649 |
@ isodefl_abs_rep_thms |
40832 | 650 |
@ IsodeflData.get (ProofContext.init_global thy) |
33845 | 651 |
in |
652 |
Goal.prove_global thy [] assms goal (fn {prems, ...} => |
|
653 |
EVERY |
|
41324
1383653efec3
make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents:
41292
diff
changeset
|
654 |
[rewrite_goals_tac (defl_apply_thms @ map_apply_thms), |
1383653efec3
make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents:
41292
diff
changeset
|
655 |
rtac (@{thm cont_parallel_fix_ind} |
1383653efec3
make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents:
41292
diff
changeset
|
656 |
OF [defl_cont_thm, map_cont_thm]) 1, |
33788 | 657 |
REPEAT (resolve_tac adm_rules 1), |
658 |
simp_tac (HOL_basic_ss addsimps bottom_rules) 1, |
|
41324
1383653efec3
make internal proofs for deflation and isodefl theorems more robust, by avoiding calls to the simplifier for beta-reduction
huffman
parents:
41292
diff
changeset
|
659 |
simp_tac (HOL_basic_ss addsimps tuple_rules) 1, |
40492 | 660 |
simp_tac (HOL_basic_ss addsimps map_ID_simps) 1, |
33788 | 661 |
REPEAT (etac @{thm conjE} 1), |
33845 | 662 |
REPEAT (resolve_tac (isodefl_rules @ prems) 1 ORELSE atac 1)]) |
40832 | 663 |
end |
664 |
val isodefl_binds = map (Binding.prefix_name "isodefl_") dbinds |
|
33789 | 665 |
fun conjuncts [] thm = [] |
666 |
| conjuncts (n::[]) thm = [(n, thm)] |
|
667 |
| conjuncts (n::ns) thm = let |
|
40832 | 668 |
val thmL = thm RS @{thm conjunct1} |
669 |
val thmR = thm RS @{thm conjunct2} |
|
670 |
in (n, thmL):: conjuncts ns thmR end |
|
33789 | 671 |
val (isodefl_thms, thy) = thy |> |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
37744
diff
changeset
|
672 |
(Global_Theory.add_thms o map (Thm.no_attributes o apsnd Drule.zero_var_indexes)) |
40832 | 673 |
(conjuncts isodefl_binds isodefl_thm) |
674 |
val thy = fold (Context.theory_map o IsodeflData.add_thm) isodefl_thms thy |
|
33793 | 675 |
|
676 |
(* prove map_ID theorems *) |
|
677 |
fun prove_map_ID_thm |
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39974
diff
changeset
|
678 |
(((map_const, (lhsT, _)), DEFL_thm), isodefl_thm) = |
33793 | 679 |
let |
40832 | 680 |
val Ts = snd (dest_Type lhsT) |
681 |
fun is_cpo T = Sign.of_sort thy (T, @{sort cpo}) |
|
682 |
val lhs = list_ccomb (map_const, map mk_ID (filter is_cpo Ts)) |
|
683 |
val goal = mk_eqs (lhs, mk_ID lhsT) |
|
33793 | 684 |
val tac = EVERY |
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39974
diff
changeset
|
685 |
[rtac @{thm isodefl_DEFL_imp_ID} 1, |
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39974
diff
changeset
|
686 |
stac DEFL_thm 1, |
33793 | 687 |
rtac isodefl_thm 1, |
40832 | 688 |
REPEAT (resolve_tac @{thms isodefl_ID_DEFL isodefl_LIFTDEFL} 1)] |
33793 | 689 |
in |
690 |
Goal.prove_global thy [] [] goal (K tac) |
|
40832 | 691 |
end |
692 |
val map_ID_binds = map (Binding.suffix_name "_map_ID") dbinds |
|
33793 | 693 |
val map_ID_thms = |
694 |
map prove_map_ID_thm |
|
40832 | 695 |
(map_consts ~~ dom_eqns ~~ DEFL_thms ~~ isodefl_thms) |
33793 | 696 |
val (_, thy) = thy |> |
40489 | 697 |
(Global_Theory.add_thms o map (rpair [Domain_Take_Proofs.map_ID_add])) |
40832 | 698 |
(map_ID_binds ~~ map_ID_thms) |
33802
48ce3a1063f2
domain_isomorphism package defines copy functions
huffman
parents:
33801
diff
changeset
|
699 |
|
35494
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
700 |
(* definitions and proofs related to take functions *) |
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
701 |
val (take_info, thy) = |
35514
a2cfa413eaab
move take-related definitions and proofs to new module; simplify map_of_typ functions
huffman
parents:
35498
diff
changeset
|
702 |
Domain_Take_Proofs.define_take_functions |
40832 | 703 |
(dbinds ~~ iso_infos) thy |
40015 | 704 |
val { take_consts, chain_take_thms, take_0_thms, take_Suc_thms, ...} = |
40832 | 705 |
take_info |
33807
ce8d2e8bca21
domain_isomorphism package defines combined copy function
huffman
parents:
33802
diff
changeset
|
706 |
|
35494
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
707 |
(* least-upper-bound lemma for take functions *) |
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
708 |
val lub_take_lemma = |
33809
033831fd9ef3
store map_ID thms in theory data; automate proofs of reach lemmas
huffman
parents:
33807
diff
changeset
|
709 |
let |
40832 | 710 |
val lhs = mk_tuple (map mk_lub take_consts) |
711 |
fun is_cpo T = Sign.of_sort thy (T, @{sort cpo}) |
|
35494
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
712 |
fun mk_map_ID (map_const, (lhsT, rhsT)) = |
40832 | 713 |
list_ccomb (map_const, map mk_ID (filter is_cpo (snd (dest_Type lhsT)))) |
714 |
val rhs = mk_tuple (map mk_map_ID (map_consts ~~ dom_eqns)) |
|
715 |
val goal = mk_trp (mk_eq (lhs, rhs)) |
|
716 |
val map_ID_thms = Domain_Take_Proofs.get_map_ID_thms thy |
|
35494
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
717 |
val start_rules = |
40771 | 718 |
@{thms lub_Pair [symmetric] ch2ch_Pair} @ chain_take_thms |
35494
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
719 |
@ @{thms pair_collapse split_def} |
40832 | 720 |
@ map_apply_thms @ map_ID_thms |
35494
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
721 |
val rules0 = |
40832 | 722 |
@{thms iterate_0 Pair_strict} @ take_0_thms |
35494
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
723 |
val rules1 = |
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
724 |
@{thms iterate_Suc Pair_fst_snd_eq fst_conv snd_conv} |
40832 | 725 |
@ take_Suc_thms |
35494
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
726 |
val tac = |
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
727 |
EVERY |
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
728 |
[simp_tac (HOL_basic_ss addsimps start_rules) 1, |
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
729 |
simp_tac (HOL_basic_ss addsimps @{thms fix_def2}) 1, |
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
730 |
rtac @{thm lub_eq} 1, |
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
731 |
rtac @{thm nat.induct} 1, |
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
732 |
simp_tac (HOL_basic_ss addsimps rules0) 1, |
40832 | 733 |
asm_full_simp_tac (beta_ss addsimps rules1) 1] |
33809
033831fd9ef3
store map_ID thms in theory data; automate proofs of reach lemmas
huffman
parents:
33807
diff
changeset
|
734 |
in |
033831fd9ef3
store map_ID thms in theory data; automate proofs of reach lemmas
huffman
parents:
33807
diff
changeset
|
735 |
Goal.prove_global thy [] [] goal (K tac) |
40832 | 736 |
end |
33809
033831fd9ef3
store map_ID thms in theory data; automate proofs of reach lemmas
huffman
parents:
33807
diff
changeset
|
737 |
|
35494
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
738 |
(* prove lub of take equals ID *) |
35773 | 739 |
fun prove_lub_take (((dbind, take_const), map_ID_thm), (lhsT, rhsT)) thy = |
33809
033831fd9ef3
store map_ID thms in theory data; automate proofs of reach lemmas
huffman
parents:
33807
diff
changeset
|
740 |
let |
40832 | 741 |
val n = Free ("n", natT) |
742 |
val goal = mk_eqs (mk_lub (lambda n (take_const $ n)), mk_ID lhsT) |
|
35494
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
743 |
val tac = |
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
744 |
EVERY |
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
745 |
[rtac @{thm trans} 1, rtac map_ID_thm 2, |
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
746 |
cut_facts_tac [lub_take_lemma] 1, |
40832 | 747 |
REPEAT (etac @{thm Pair_inject} 1), atac 1] |
748 |
val lub_take_thm = Goal.prove_global thy [] [] goal (K tac) |
|
33809
033831fd9ef3
store map_ID thms in theory data; automate proofs of reach lemmas
huffman
parents:
33807
diff
changeset
|
749 |
in |
35773 | 750 |
add_qualified_thm "lub_take" (dbind, lub_take_thm) thy |
40832 | 751 |
end |
35494
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
752 |
val (lub_take_thms, thy) = |
45c9a8278faf
domain package no longer generates copy functions; all proofs use take functions instead
huffman
parents:
35490
diff
changeset
|
753 |
fold_map prove_lub_take |
40832 | 754 |
(dbinds ~~ take_consts ~~ map_ID_thms ~~ dom_eqns) thy |
33809
033831fd9ef3
store map_ID thms in theory data; automate proofs of reach lemmas
huffman
parents:
33807
diff
changeset
|
755 |
|
35654
7a15e181bf3b
move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents:
35640
diff
changeset
|
756 |
(* prove additional take theorems *) |
7a15e181bf3b
move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents:
35640
diff
changeset
|
757 |
val (take_info2, thy) = |
7a15e181bf3b
move proofs of reach and take lemmas to domain_take_proofs.ML
huffman
parents:
35640
diff
changeset
|
758 |
Domain_Take_Proofs.add_lub_take_theorems |
40832 | 759 |
(dbinds ~~ iso_infos) take_info lub_take_thms thy |
33774 | 760 |
in |
35659 | 761 |
((iso_infos, take_info2), thy) |
40832 | 762 |
end |
33774 | 763 |
|
40832 | 764 |
val domain_isomorphism = gen_domain_isomorphism cert_typ |
765 |
val domain_isomorphism_cmd = snd oo gen_domain_isomorphism read_typ |
|
33774 | 766 |
|
767 |
(******************************************************************************) |
|
768 |
(******************************** outer syntax ********************************) |
|
769 |
(******************************************************************************) |
|
770 |
||
771 |
local |
|
772 |
||
34149
a0efb4754cb7
add 'morphisms' option to domain_isomorphism command
huffman
parents:
33971
diff
changeset
|
773 |
val parse_domain_iso : |
a0efb4754cb7
add 'morphisms' option to domain_isomorphism command
huffman
parents:
33971
diff
changeset
|
774 |
(string list * binding * mixfix * string * (binding * binding) option) |
a0efb4754cb7
add 'morphisms' option to domain_isomorphism command
huffman
parents:
33971
diff
changeset
|
775 |
parser = |
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
776 |
(Parse.type_args -- Parse.binding -- Parse.opt_mixfix -- (Parse.$$$ "=" |-- Parse.typ) -- |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
777 |
Scan.option (Parse.$$$ "morphisms" |-- Parse.!!! (Parse.binding -- Parse.binding))) |
40832 | 778 |
>> (fn ((((vs, t), mx), rhs), morphs) => (vs, t, mx, rhs, morphs)) |
33774 | 779 |
|
40832 | 780 |
val parse_domain_isos = Parse.and_list1 parse_domain_iso |
33774 | 781 |
|
782 |
in |
|
783 |
||
784 |
val _ = |
|
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
785 |
Outer_Syntax.command "domain_isomorphism" "define domain isomorphisms (HOLCF)" |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36610
diff
changeset
|
786 |
Keyword.thy_decl |
40832 | 787 |
(parse_domain_isos >> (Toplevel.theory o domain_isomorphism_cmd)) |
33774 | 788 |
|
40832 | 789 |
end |
33774 | 790 |
|
40832 | 791 |
end |