author | blanchet |
Thu, 06 Mar 2014 15:40:33 +0100 | |
changeset 55945 | e96383acecf9 |
parent 54742 | 7a86358a3c0b |
child 56519 | c1048f5bbb45 |
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 |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
12 |
|
45598 | 13 |
val quotient_def: |
14 |
(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
|
15 |
local_theory -> Proof.state |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
16 |
|
45598 | 17 |
val quotient_def_cmd: |
18 |
(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
|
19 |
local_theory -> Proof.state |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
20 |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
21 |
end; |
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 |
structure Quotient_Def: QUOTIENT_DEF = |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
24 |
struct |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
25 |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
26 |
(** Interface and Syntax Setup **) |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
27 |
|
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 |
(* 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
|
29 |
|
47698 | 30 |
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
|
31 |
|
47698 | 32 |
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
|
33 |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
34 |
(* The ML-interface for a quotient definition takes |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
35 |
as argument: |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
36 |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
37 |
- an optional binding and mixfix annotation |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
38 |
- attributes |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
39 |
- the new constant as term |
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
40 |
- 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
|
41 |
- respectfulness theorem for the rhs |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
42 |
|
45279 | 43 |
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
|
44 |
|
41444 | 45 |
Restriction: At the moment the left- and right-hand |
46 |
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
|
47 |
*) |
41444 | 48 |
fun error_msg bind str = |
49 |
let |
|
50 |
val name = Binding.name_of bind |
|
48992 | 51 |
val pos = Position.here (Binding.pos_of bind) |
41444 | 52 |
in |
53 |
error ("Head of quotient_definition " ^ |
|
54 |
quote str ^ " differs from declaration " ^ name ^ pos) |
|
55 |
end |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
56 |
|
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
57 |
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
|
58 |
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
|
59 |
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
|
60 |
val qty = fastype_of lhs |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
61 |
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
|
62 |
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
|
63 |
val prop = Syntax.check_term lthy (Logic.mk_equals (lhs, absrep_trm)) |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
64 |
val (_, prop') = Local_Defs.cert_def lthy prop |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
65 |
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
|
66 |
|
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
|
67 |
val ((trm, (_ , def_thm)), lthy') = |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
68 |
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
|
69 |
|
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
70 |
(* data storage *) |
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
|
71 |
val qconst_data = {qconst = trm, rconst = rhs, def = def_thm} |
47198
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 |
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
|
74 |
|> Binding.qualify true defname |
cfd8ff62eab1
use qualified names for rsp and rep_eq theorems in quotient_def
kuncar
parents:
47181
diff
changeset
|
75 |
|
cfd8ff62eab1
use qualified names for rsp and rep_eq theorems in quotient_def
kuncar
parents:
47181
diff
changeset
|
76 |
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
|
77 |
val rsp_thm_name = qualify lhs_name "rsp" |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
78 |
|
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
79 |
val lthy'' = lthy' |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
80 |
|> 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
|
81 |
(fn phi => |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
82 |
(case Quotient_Info.transform_quotconsts phi qconst_data of |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
83 |
qcinfo as {qconst = Const (c, _), ...} => |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
84 |
Quotient_Info.update_quotconsts c qcinfo |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
85 |
| _ => I)) |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
86 |
|> (snd oo Local_Theory.note) |
47156 | 87 |
((rsp_thm_name, [Attrib.internal (K Quotient_Info.rsp_rules_add)]), |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
88 |
[rsp_thm]) |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
89 |
in |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
90 |
(qconst_data, lthy'') |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
91 |
end |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
92 |
|
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
93 |
fun mk_readable_rsp_thm_eq tm lthy = |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
94 |
let |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
95 |
val ctm = cterm_of (Proof_Context.theory_of lthy) tm |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
96 |
|
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
97 |
fun norm_fun_eq ctm = |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
98 |
let |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
99 |
fun abs_conv2 cv = Conv.abs_conv (K (Conv.abs_conv (K cv) lthy)) lthy |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
100 |
fun erase_quants ctm' = |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
101 |
case (Thm.term_of ctm') of |
47181
b351ad77eb78
some tuning while reviewing the current state of the quotient_def package
bulwahn
parents:
47156
diff
changeset
|
102 |
Const (@{const_name HOL.eq}, _) $ _ $ _ => Conv.all_conv ctm' |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
103 |
| _ => (Conv.binder_conv (K erase_quants) lthy then_conv |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
104 |
Conv.rewr_conv @{thm fun_eq_iff[symmetric, THEN eq_reflection]}) ctm' |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
105 |
in |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
106 |
(abs_conv2 erase_quants then_conv Thm.eta_conversion) ctm |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
107 |
end |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
108 |
|
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
109 |
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
|
110 |
let |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
111 |
val unfold_conv = Conv.rewrs_conv |
55945 | 112 |
[@{thm rel_fun_eq_invariant[THEN eq_reflection]}, @{thm rel_fun_eq_rel[THEN eq_reflection]}, |
113 |
@{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
|
114 |
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
|
115 |
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
|
116 |
in |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
117 |
case (Thm.term_of ctm) of |
55945 | 118 |
Const (@{const_name rel_fun}, _) $ _ $ _ => |
47181
b351ad77eb78
some tuning while reviewing the current state of the quotient_def package
bulwahn
parents:
47156
diff
changeset
|
119 |
(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
|
120 |
| _ => 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
|
121 |
end |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
122 |
|
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
|
123 |
val unfold_ret_val_invs = Conv.bottom_conv |
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
|
124 |
(K (Conv.try_conv (Conv.rewr_conv @{thm invariant_same_args}))) lthy |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
125 |
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
|
126 |
val univq_conv = Conv.rewr_conv @{thm HOL.all_simps(6)[symmetric, THEN eq_reflection]} |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
127 |
val univq_prenex_conv = Conv.top_conv (K (Conv.try_conv univq_conv)) lthy |
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
|
128 |
val beta_conv = Thm.beta_conversion true |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
129 |
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
|
130 |
(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
|
131 |
in |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
50902
diff
changeset
|
132 |
Object_Logic.rulify lthy (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
|
133 |
end |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
134 |
|
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
135 |
|
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
136 |
|
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
137 |
fun gen_quotient_def prep_vars prep_term (raw_var, (attr, (lhs_raw, rhs_raw))) lthy = |
41444 | 138 |
let |
43608
294570668f25
parse term in auxiliary context augmented with variable;
krauss
parents:
42494
diff
changeset
|
139 |
val (vars, ctxt) = prep_vars (the_list raw_var) lthy |
43663
e8c80bbc0c5d
re-check to explicitly propagate a given type constraint to lhs -- necessary to trigger type improvement in an instantiation target
krauss
parents:
43608
diff
changeset
|
140 |
val T_opt = (case vars of [(_, SOME T, _)] => SOME T | _ => NONE) |
e8c80bbc0c5d
re-check to explicitly propagate a given type constraint to lhs -- necessary to trigger type improvement in an instantiation target
krauss
parents:
43608
diff
changeset
|
141 |
val lhs = prep_term T_opt ctxt lhs_raw |
e8c80bbc0c5d
re-check to explicitly propagate a given type constraint to lhs -- necessary to trigger type improvement in an instantiation target
krauss
parents:
43608
diff
changeset
|
142 |
val rhs = prep_term NONE ctxt rhs_raw |
43608
294570668f25
parse term in auxiliary context augmented with variable;
krauss
parents:
42494
diff
changeset
|
143 |
|
41444 | 144 |
val (lhs_str, lhs_ty) = dest_Free lhs handle TERM _ => error "Constant already defined." |
145 |
val _ = if null (strip_abs_vars rhs) then () else error "The definiens cannot be an abstraction" |
|
146 |
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
|
147 |
|
43608
294570668f25
parse term in auxiliary context augmented with variable;
krauss
parents:
42494
diff
changeset
|
148 |
val var = |
294570668f25
parse term in auxiliary context augmented with variable;
krauss
parents:
42494
diff
changeset
|
149 |
(case vars of |
294570668f25
parse term in auxiliary context augmented with variable;
krauss
parents:
42494
diff
changeset
|
150 |
[] => (Binding.name lhs_str, NoSyn) |
294570668f25
parse term in auxiliary context augmented with variable;
krauss
parents:
42494
diff
changeset
|
151 |
| [(binding, _, mx)] => |
294570668f25
parse term in auxiliary context augmented with variable;
krauss
parents:
42494
diff
changeset
|
152 |
if Variable.check_name binding = lhs_str then (binding, mx) |
294570668f25
parse term in auxiliary context augmented with variable;
krauss
parents:
42494
diff
changeset
|
153 |
else error_msg binding lhs_str |
294570668f25
parse term in auxiliary context augmented with variable;
krauss
parents:
42494
diff
changeset
|
154 |
| _ => raise Match) |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
155 |
|
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
156 |
fun try_to_prove_refl thm = |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
157 |
let |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
158 |
val lhs_eq = |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
159 |
thm |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
160 |
|> prop_of |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
161 |
|> Logic.dest_implies |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
162 |
|> fst |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
163 |
|> strip_all_body |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
164 |
|> 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
|
165 |
in |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
166 |
case lhs_eq of |
47181
b351ad77eb78
some tuning while reviewing the current state of the quotient_def package
bulwahn
parents:
47156
diff
changeset
|
167 |
SOME (Const (@{const_name HOL.eq}, _) $ _ $ _) => 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
|
168 |
| 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
|
169 |
Type (typ_name, _) => |
cb2b940e2fdf
avoid handling arbitrary exceptions, notably physical interrupts that would make the program erratic;
wenzelm
parents:
48992
diff
changeset
|
170 |
try (fn () => |
cb2b940e2fdf
avoid handling arbitrary exceptions, notably physical interrupts that would make the program erratic;
wenzelm
parents:
48992
diff
changeset
|
171 |
#equiv_thm (the (Quotient_Info.lookup_quotients lthy typ_name)) |
cb2b940e2fdf
avoid handling arbitrary exceptions, notably physical interrupts that would make the program erratic;
wenzelm
parents:
48992
diff
changeset
|
172 |
RS @{thm Equiv_Relations.equivp_reflp} RS thm) () |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
173 |
| _ => NONE |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
174 |
) |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
175 |
| _ => NONE |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
176 |
end |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
177 |
|
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
178 |
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
|
179 |
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
|
180 |
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
|
181 |
val maybe_proven_rsp_thm = try_to_prove_refl readable_rsp_thm_eq |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
182 |
val (readable_rsp_tm, _) = Logic.dest_implies (prop_of readable_rsp_thm_eq) |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
183 |
|
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
184 |
fun after_qed thm_list lthy = |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
185 |
let |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
186 |
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
|
187 |
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
|
188 |
[] => the maybe_proven_rsp_thm |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
189 |
| [[thm]] => Goal.prove ctxt [] [] internal_rsp_tm |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
50902
diff
changeset
|
190 |
(fn _ => rtac readable_rsp_thm_eq 1 THEN 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
|
191 |
in |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
192 |
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
|
193 |
end |
37530
70d03844b2f9
export of proper information in the ML-interface of the quotient package
Christian Urban <urbanc@in.tum.de>
parents:
36960
diff
changeset
|
194 |
|
41444 | 195 |
in |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
196 |
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
|
197 |
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
|
198 |
| NONE => Proof.theorem NONE after_qed [[(readable_rsp_tm,[])]] lthy |
41444 | 199 |
end |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
200 |
|
43663
e8c80bbc0c5d
re-check to explicitly propagate a given type constraint to lhs -- necessary to trigger type improvement in an instantiation target
krauss
parents:
43608
diff
changeset
|
201 |
fun check_term' cnstr ctxt = |
e8c80bbc0c5d
re-check to explicitly propagate a given type constraint to lhs -- necessary to trigger type improvement in an instantiation target
krauss
parents:
43608
diff
changeset
|
202 |
Syntax.check_term ctxt o (case cnstr of SOME T => Type.constraint T | _ => I) |
45598 | 203 |
|
204 |
fun read_term' cnstr ctxt = |
|
45929
d7d6c8cfb6f6
removing some debug output in quotient_definition
bulwahn
parents:
45826
diff
changeset
|
205 |
check_term' cnstr ctxt o Syntax.parse_term ctxt |
43663
e8c80bbc0c5d
re-check to explicitly propagate a given type constraint to lhs -- necessary to trigger type improvement in an instantiation target
krauss
parents:
43608
diff
changeset
|
206 |
|
e8c80bbc0c5d
re-check to explicitly propagate a given type constraint to lhs -- necessary to trigger type improvement in an instantiation target
krauss
parents:
43608
diff
changeset
|
207 |
val quotient_def = gen_quotient_def Proof_Context.cert_vars check_term' |
45598 | 208 |
val quotient_def_cmd = gen_quotient_def Proof_Context.read_vars read_term' |
43608
294570668f25
parse term in auxiliary context augmented with variable;
krauss
parents:
42494
diff
changeset
|
209 |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
210 |
|
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
211 |
(* parser and command *) |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
212 |
val quotdef_parser = |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
213 |
Scan.option Parse_Spec.constdecl -- |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
214 |
Parse.!!! (Parse_Spec.opt_thm_name ":" -- (Parse.term --| @{keyword "is"} -- Parse.term)) |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
215 |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
216 |
val _ = |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
217 |
Outer_Syntax.local_theory_to_proof @{command_spec "quotient_definition"} |
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
218 |
"definition for constants over the quotient type" |
47091
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
219 |
(quotdef_parser >> quotient_def_cmd) |
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
kuncar
parents:
46961
diff
changeset
|
220 |
|
35222
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
221 |
|
4f1fba00f66d
Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
222 |
end; (* structure *) |