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