author | wenzelm |
Sun, 07 Nov 2021 19:53:37 +0100 | |
changeset 74726 | 33ed2eb06d68 |
parent 74535 | 2f8e8dc9b522 |
child 78095 | bc42c074e58f |
permissions | -rw-r--r-- |
37744 | 1 |
(* Title: HOL/Tools/Quotient/quotient_def.ML |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
2 |
Author: Cezary Kaliszyk and Christian Urban |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
3 |
|
35788 | 4 |
Definitions for constants on quotient types. |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
5 |
*) |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
6 |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
7 |
signature QUOTIENT_DEF = |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
8 |
sig |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
9 |
val add_quotient_def: |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
10 |
((binding * mixfix) * Attrib.binding) * (term * term) -> thm -> |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
11 |
local_theory -> Quotient_Info.quotconsts * local_theory |
45598 | 12 |
val quotient_def: |
13 |
(binding * typ option * mixfix) option * (Attrib.binding * (term * term)) -> |
|
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
14 |
local_theory -> Proof.state |
45598 | 15 |
val quotient_def_cmd: |
16 |
(binding * string option * mixfix) option * (Attrib.binding * (string * string)) -> |
|
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
17 |
local_theory -> Proof.state |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
18 |
end; |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
19 |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
20 |
structure Quotient_Def: QUOTIENT_DEF = |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
21 |
struct |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
22 |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
23 |
(** Interface and Syntax Setup **) |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
24 |
|
47096
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents:
47091
diff
changeset
|
25 |
(* Generation of the code certificate from the rsp theorem *) |
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents:
47091
diff
changeset
|
26 |
|
47698 | 27 |
open Lifting_Util |
47096
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents:
47091
diff
changeset
|
28 |
|
47698 | 29 |
infix 0 MRSL |
47096
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents:
47091
diff
changeset
|
30 |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
31 |
(* The ML-interface for a quotient definition takes |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
32 |
as argument: |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
33 |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
34 |
- an optional binding and mixfix annotation |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
35 |
- attributes |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
36 |
- the new constant as term |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
37 |
- the rhs of the definition as term |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
38 |
- respectfulness theorem for the rhs |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
39 |
|
45279 | 40 |
It stores the qconst_info in the quotconsts data slot. |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
41 |
|
41444 | 42 |
Restriction: At the moment the left- and right-hand |
43 |
side of the definition must be a constant. |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
44 |
*) |
41444 | 45 |
fun error_msg bind str = |
46 |
let |
|
47 |
val name = Binding.name_of bind |
|
48992 | 48 |
val pos = Position.here (Binding.pos_of bind) |
41444 | 49 |
in |
50 |
error ("Head of quotient_definition " ^ |
|
51 |
quote str ^ " differs from declaration " ^ name ^ pos) |
|
52 |
end |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
53 |
|
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
54 |
fun add_quotient_def ((var, (name, atts)), (lhs, rhs)) rsp_thm lthy = |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
55 |
let |
47096
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents:
47091
diff
changeset
|
56 |
val rty = fastype_of rhs |
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents:
47091
diff
changeset
|
57 |
val qty = fastype_of lhs |
60379 | 58 |
val absrep_trm = |
47096
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents:
47091
diff
changeset
|
59 |
Quotient_Term.absrep_fun lthy Quotient_Term.AbsF (rty, qty) $ rhs |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
60 |
val prop = Syntax.check_term lthy (Logic.mk_equals (lhs, absrep_trm)) |
63395 | 61 |
val (_, prop') = Local_Defs.cert_def lthy (K []) prop |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
62 |
val (_, newrhs) = Local_Defs.abs_def prop' |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
63 |
|
67632 | 64 |
val ((qconst, (_ , def)), lthy') = |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
65 |
Local_Theory.define (var, ((Thm.def_binding_optional (#1 var) name, atts), newrhs)) lthy |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
66 |
|
67632 | 67 |
fun qconst_data phi = |
68 |
Quotient_Info.transform_quotconsts phi {qconst = qconst, rconst = rhs, def = def} |
|
60379 | 69 |
|
47198
cfd8ff62eab1
use qualified names for rsp and rep_eq theorems in quotient_def
kuncar
parents:
47181
diff
changeset
|
70 |
fun qualify defname suffix = Binding.name suffix |
cfd8ff62eab1
use qualified names for rsp and rep_eq theorems in quotient_def
kuncar
parents:
47181
diff
changeset
|
71 |
|> Binding.qualify true defname |
cfd8ff62eab1
use qualified names for rsp and rep_eq theorems in quotient_def
kuncar
parents:
47181
diff
changeset
|
72 |
|
cfd8ff62eab1
use qualified names for rsp and rep_eq theorems in quotient_def
kuncar
parents:
47181
diff
changeset
|
73 |
val lhs_name = Binding.name_of (#1 var) |
cfd8ff62eab1
use qualified names for rsp and rep_eq theorems in quotient_def
kuncar
parents:
47181
diff
changeset
|
74 |
val rsp_thm_name = qualify lhs_name "rsp" |
60379 | 75 |
|
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
76 |
val lthy'' = lthy' |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
77 |
|> Local_Theory.declaration {syntax = false, pervasive = true} |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
78 |
(fn phi => |
67632 | 79 |
(case qconst_data phi of |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
80 |
qcinfo as {qconst = Const (c, _), ...} => |
59157 | 81 |
Quotient_Info.update_quotconsts (c, qcinfo) |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
82 |
| _ => I)) |
60379 | 83 |
|> (snd oo Local_Theory.note) |
57960 | 84 |
((rsp_thm_name, @{attributes [quot_respect]}), [rsp_thm]) |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
85 |
in |
67632 | 86 |
(qconst_data Morphism.identity, lthy'') |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
87 |
end |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
88 |
|
74535 | 89 |
fun mk_readable_rsp_thm_eq tm ctxt = |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
90 |
let |
74535 | 91 |
val ctm = Thm.cterm_of ctxt tm |
60379 | 92 |
|
74535 | 93 |
fun abs_conv2 cv = Conv.abs_conv (Conv.abs_conv (cv o #2) o #2) ctxt |
94 |
fun erase_quants ctxt' ctm' = |
|
95 |
case (Thm.term_of ctm') of |
|
96 |
Const (\<^const_name>\<open>HOL.eq\<close>, _) $ _ $ _ => Conv.all_conv ctm' |
|
97 |
| _ => (Conv.binder_conv (erase_quants o #2) ctxt' then_conv |
|
98 |
Conv.rewr_conv @{thm fun_eq_iff[symmetric, THEN eq_reflection]}) ctm' |
|
99 |
val norm_fun_eq = abs_conv2 erase_quants then_conv Thm.eta_conversion |
|
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
100 |
|
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
101 |
fun simp_arrows_conv ctm = |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
102 |
let |
60379 | 103 |
val unfold_conv = Conv.rewrs_conv |
104 |
[@{thm rel_fun_eq_eq_onp[THEN eq_reflection]}, @{thm rel_fun_eq_rel[THEN eq_reflection]}, |
|
55945 | 105 |
@{thm rel_fun_def[THEN eq_reflection]}] |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
106 |
val left_conv = simp_arrows_conv then_conv Conv.try_conv norm_fun_eq |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
107 |
fun binop_conv2 cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2 |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
108 |
in |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
109 |
case (Thm.term_of ctm) of |
69593 | 110 |
Const (\<^const_name>\<open>rel_fun\<close>, _) $ _ $ _ => |
47181
b351ad77eb78
some tuning while reviewing the current state of the quotient_def package
bulwahn
parents:
47156
diff
changeset
|
111 |
(binop_conv2 left_conv simp_arrows_conv then_conv unfold_conv) ctm |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
112 |
| _ => Conv.all_conv ctm |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
113 |
end |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
114 |
|
60379 | 115 |
val unfold_ret_val_invs = Conv.bottom_conv |
74535 | 116 |
(K (Conv.try_conv (Conv.rewr_conv @{thm eq_onp_same_args[THEN eq_reflection]}))) ctxt |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
117 |
val simp_conv = Conv.arg_conv (Conv.fun2_conv simp_arrows_conv) |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
118 |
val univq_conv = Conv.rewr_conv @{thm HOL.all_simps(6)[symmetric, THEN eq_reflection]} |
74535 | 119 |
val univq_prenex_conv = Conv.top_conv (K (Conv.try_conv univq_conv)) ctxt |
47096
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents:
47091
diff
changeset
|
120 |
val beta_conv = Thm.beta_conversion true |
60379 | 121 |
val eq_thm = |
47096
3ea48c19673e
generation of a code certificate from a respectfulness theorem for constants lifted by the quotient_definition command & setup_lifting command: setups Quotient infrastructure from a typedef theorem
kuncar
parents:
47091
diff
changeset
|
122 |
(simp_conv then_conv univq_prenex_conv then_conv beta_conv then_conv unfold_ret_val_invs) ctm |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
123 |
in |
74535 | 124 |
Object_Logic.rulify ctxt (eq_thm RS Drule.equal_elim_rule2) |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
125 |
end |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
126 |
|
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
127 |
|
60379 | 128 |
fun gen_quotient_def prep_var parse_term (raw_var, (attr, (raw_lhs, raw_rhs))) lthy = |
41444 | 129 |
let |
60379 | 130 |
val (opt_var, ctxt) = |
131 |
(case raw_var of |
|
132 |
NONE => (NONE, lthy) |
|
62956 | 133 |
| SOME var => prep_var var lthy |>> SOME) |
134 |
val lhs_constraints = (case opt_var of SOME (_, SOME T, _) => [T] | _ => []) |
|
60379 | 135 |
|
62956 | 136 |
fun prep_term Ts = parse_term ctxt #> fold Type.constraint Ts #> Syntax.check_term ctxt; |
137 |
val lhs = prep_term lhs_constraints raw_lhs |
|
138 |
val rhs = prep_term [] raw_rhs |
|
43608
294570668f25
parse term in auxiliary context augmented with variable;
krauss
parents:
42494
diff
changeset
|
139 |
|
62953 | 140 |
val (lhs_str, lhs_ty) = dest_Free lhs handle TERM _ => error "Constant already defined" |
41444 | 141 |
val _ = if null (strip_abs_vars rhs) then () else error "The definiens cannot be an abstraction" |
142 |
val _ = if is_Const rhs then () else warning "The definiens is not a constant" |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
143 |
|
43608
294570668f25
parse term in auxiliary context augmented with variable;
krauss
parents:
42494
diff
changeset
|
144 |
val var = |
60379 | 145 |
(case opt_var of |
146 |
NONE => (Binding.name lhs_str, NoSyn) |
|
147 |
| SOME (binding, _, mx) => |
|
43608
294570668f25
parse term in auxiliary context augmented with variable;
krauss
parents:
42494
diff
changeset
|
148 |
if Variable.check_name binding = lhs_str then (binding, mx) |
60379 | 149 |
else error_msg binding lhs_str); |
150 |
||
151 |
fun try_to_prove_refl thm = |
|
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
152 |
let |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
153 |
val lhs_eq = |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
154 |
thm |
59582 | 155 |
|> Thm.prop_of |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
156 |
|> Logic.dest_implies |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
157 |
|> fst |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
158 |
|> strip_all_body |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
159 |
|> try HOLogic.dest_Trueprop |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
160 |
in |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
161 |
case lhs_eq of |
69593 | 162 |
SOME (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ _ $ _) => SOME (@{thm refl} RS thm) |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
163 |
| SOME _ => (case body_type (fastype_of lhs) of |
50902
cb2b940e2fdf
avoid handling arbitrary exceptions, notably physical interrupts that would make the program erratic;
wenzelm
parents:
48992
diff
changeset
|
164 |
Type (typ_name, _) => |
73551 | 165 |
\<^try>\<open> |
60379 | 166 |
#equiv_thm (the (Quotient_Info.lookup_quotients lthy typ_name)) |
73551 | 167 |
RS @{thm Equiv_Relations.equivp_reflp} RS thm\<close> |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
168 |
| _ => NONE |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
169 |
) |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
170 |
| _ => NONE |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
171 |
end |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
172 |
|
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
173 |
val rsp_rel = Quotient_Term.equiv_relation lthy (fastype_of rhs, lhs_ty) |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
174 |
val internal_rsp_tm = HOLogic.mk_Trueprop (Syntax.check_term lthy (rsp_rel $ rhs $ rhs)) |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
175 |
val readable_rsp_thm_eq = mk_readable_rsp_thm_eq internal_rsp_tm lthy |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
176 |
val maybe_proven_rsp_thm = try_to_prove_refl readable_rsp_thm_eq |
59582 | 177 |
val (readable_rsp_tm, _) = Logic.dest_implies (Thm.prop_of readable_rsp_thm_eq) |
60379 | 178 |
|
179 |
fun after_qed thm_list lthy = |
|
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
180 |
let |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
181 |
val internal_rsp_thm = |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
182 |
case thm_list of |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
183 |
[] => the maybe_proven_rsp_thm |
60379 | 184 |
| [[thm]] => Goal.prove ctxt [] [] internal_rsp_tm |
60752 | 185 |
(fn _ => |
186 |
resolve_tac ctxt [readable_rsp_thm_eq] 1 THEN |
|
187 |
Proof_Context.fact_tac ctxt [thm] 1) |
|
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
188 |
in |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
189 |
snd (add_quotient_def ((var, attr), (lhs, rhs)) internal_rsp_thm lthy) |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
190 |
end |
41444 | 191 |
in |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
192 |
case maybe_proven_rsp_thm of |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
193 |
SOME _ => Proof.theorem NONE after_qed [] lthy |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
194 |
| NONE => Proof.theorem NONE after_qed [[(readable_rsp_tm,[])]] lthy |
41444 | 195 |
end |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
196 |
|
60379 | 197 |
val quotient_def = gen_quotient_def Proof_Context.cert_var (K I) |
198 |
val quotient_def_cmd = gen_quotient_def Proof_Context.read_var Syntax.parse_term |
|
43608
294570668f25
parse term in auxiliary context augmented with variable;
krauss
parents:
42494
diff
changeset
|
199 |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
200 |
|
60669 | 201 |
(* command syntax *) |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
202 |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
203 |
val _ = |
69593 | 204 |
Outer_Syntax.local_theory_to_proof \<^command_keyword>\<open>quotient_definition\<close> |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
205 |
"definition for constants over the quotient type" |
60669 | 206 |
(Scan.option Parse_Spec.constdecl -- |
69593 | 207 |
Parse.!!! (Parse_Spec.opt_thm_name ":" -- (Parse.term -- (\<^keyword>\<open>is\<close> |-- Parse.term))) |
60669 | 208 |
>> quotient_def_cmd); |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
209 |
|
60669 | 210 |
end; |