| author | wenzelm | 
| Sat, 18 Jul 2015 20:47:08 +0200 | |
| changeset 60752 | b48830b670a1 | 
| parent 60669 | 0e745bd11c55 | 
| child 62775 | b486f512a471 | 
| 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))  | 
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
61  | 
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
 | 
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  | 
|
| 
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
 | 
64  | 
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
 | 
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  | 
|
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
67  | 
(* 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
 | 
68  | 
    val qconst_data = {qconst = trm, rconst = rhs, def = def_thm}
 | 
| 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 =>  | 
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
79  | 
(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
 | 
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  | 
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
86  | 
(qconst_data, lthy'')  | 
| 
 
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  | 
|
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
89  | 
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
 | 
90  | 
let  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59582 
diff
changeset
 | 
91  | 
val ctm = Thm.cterm_of lthy tm  | 
| 60379 | 92  | 
|
93  | 
fun norm_fun_eq ctm =  | 
|
| 
47091
 
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  | 
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
 | 
96  | 
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
 | 
97  | 
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
 | 
98  | 
            Const (@{const_name HOL.eq}, _) $ _ $ _ => Conv.all_conv ctm'
 | 
| 60379 | 99  | 
| _ => (Conv.binder_conv (K erase_quants) lthy then_conv  | 
| 
47091
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
100  | 
              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
 | 
101  | 
in  | 
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
102  | 
(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
 | 
103  | 
end  | 
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
104  | 
|
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
105  | 
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
 | 
106  | 
let  | 
| 60379 | 107  | 
val unfold_conv = Conv.rewrs_conv  | 
108  | 
          [@{thm rel_fun_eq_eq_onp[THEN eq_reflection]}, @{thm rel_fun_eq_rel[THEN eq_reflection]},
 | 
|
| 55945 | 109  | 
            @{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
 | 
110  | 
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
 | 
111  | 
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
 | 
112  | 
in  | 
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
113  | 
case (Thm.term_of ctm) of  | 
| 60379 | 114  | 
          Const (@{const_name rel_fun}, _) $ _ $ _ =>
 | 
| 
47181
 
b351ad77eb78
some tuning while reviewing the current state of the quotient_def package
 
bulwahn 
parents: 
47156 
diff
changeset
 | 
115  | 
(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
 | 
116  | 
| _ => 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
 | 
117  | 
end  | 
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
118  | 
|
| 60379 | 119  | 
val unfold_ret_val_invs = Conv.bottom_conv  | 
120  | 
      (K (Conv.try_conv (Conv.rewr_conv @{thm eq_onp_same_args[THEN eq_reflection]}))) lthy
 | 
|
| 
47091
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
121  | 
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
 | 
122  | 
    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
 | 
123  | 
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
 | 
124  | 
val beta_conv = Thm.beta_conversion true  | 
| 60379 | 125  | 
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
 | 
126  | 
(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
 | 
127  | 
in  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
50902 
diff
changeset
 | 
128  | 
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
 | 
129  | 
end  | 
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
130  | 
|
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
131  | 
|
| 60379 | 132  | 
fun gen_quotient_def prep_var parse_term (raw_var, (attr, (raw_lhs, raw_rhs))) lthy =  | 
| 41444 | 133  | 
let  | 
| 60379 | 134  | 
val (opt_var, ctxt) =  | 
135  | 
(case raw_var of  | 
|
136  | 
NONE => (NONE, lthy)  | 
|
137  | 
| SOME var => prep_var var lthy |>> SOME)  | 
|
138  | 
val lhs_constraint = (case opt_var of SOME (_, SOME T, _) => T | _ => dummyT)  | 
|
139  | 
||
140  | 
fun prep_term T = parse_term ctxt #> Type.constraint T #> Syntax.check_term ctxt;  | 
|
141  | 
val lhs = prep_term lhs_constraint raw_lhs  | 
|
142  | 
val rhs = prep_term dummyT raw_rhs  | 
|
| 
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 =  | 
| 60379 | 149  | 
(case opt_var of  | 
150  | 
NONE => (Binding.name lhs_str, NoSyn)  | 
|
151  | 
| SOME (binding, _, mx) =>  | 
|
| 
43608
 
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)  | 
| 60379 | 153  | 
else error_msg binding lhs_str);  | 
154  | 
||
155  | 
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
 | 
156  | 
let  | 
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
157  | 
val lhs_eq =  | 
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
158  | 
thm  | 
| 59582 | 159  | 
|> 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
 | 
160  | 
|> Logic.dest_implies  | 
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
161  | 
|> fst  | 
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
162  | 
|> strip_all_body  | 
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
163  | 
|> 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
 | 
164  | 
in  | 
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
165  | 
case lhs_eq of  | 
| 
47181
 
b351ad77eb78
some tuning while reviewing the current state of the quotient_def package
 
bulwahn 
parents: 
47156 
diff
changeset
 | 
166  | 
          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
 | 
167  | 
| 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
 | 
168  | 
Type (typ_name, _) =>  | 
| 
 
cb2b940e2fdf
avoid handling arbitrary exceptions, notably physical interrupts that would make the program erratic;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
169  | 
try (fn () =>  | 
| 60379 | 170  | 
#equiv_thm (the (Quotient_Info.lookup_quotients lthy typ_name))  | 
| 
50902
 
cb2b940e2fdf
avoid handling arbitrary exceptions, notably physical interrupts that would make the program erratic;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
171  | 
                  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
 | 
172  | 
| _ => NONE  | 
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
173  | 
)  | 
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
174  | 
| _ => NONE  | 
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
175  | 
end  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
176  | 
|
| 
47091
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
177  | 
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
 | 
178  | 
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
 | 
179  | 
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
 | 
180  | 
val maybe_proven_rsp_thm = try_to_prove_refl readable_rsp_thm_eq  | 
| 59582 | 181  | 
val (readable_rsp_tm, _) = Logic.dest_implies (Thm.prop_of readable_rsp_thm_eq)  | 
| 60379 | 182  | 
|
183  | 
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
 | 
184  | 
let  | 
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
185  | 
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
 | 
186  | 
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
 | 
187  | 
[] => the maybe_proven_rsp_thm  | 
| 60379 | 188  | 
| [[thm]] => Goal.prove ctxt [] [] internal_rsp_tm  | 
| 60752 | 189  | 
(fn _ =>  | 
190  | 
resolve_tac ctxt [readable_rsp_thm_eq] 1 THEN  | 
|
191  | 
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
 | 
192  | 
in  | 
| 
 
d5cd13aca90b
respectfulness theorem has to be proved if a new constant is lifted by quotient_definition
 
kuncar 
parents: 
46961 
diff
changeset
 | 
193  | 
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
 | 
194  | 
end  | 
| 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  | 
|
| 60379 | 201  | 
val quotient_def = gen_quotient_def Proof_Context.cert_var (K I)  | 
202  | 
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
 | 
203  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
204  | 
|
| 60669 | 205  | 
(* command syntax *)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
206  | 
|
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
207  | 
val _ =  | 
| 
59936
 
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
 
wenzelm 
parents: 
59621 
diff
changeset
 | 
208  | 
  Outer_Syntax.local_theory_to_proof @{command_keyword quotient_definition}
 | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
209  | 
"definition for constants over the quotient type"  | 
| 60669 | 210  | 
(Scan.option Parse_Spec.constdecl --  | 
211  | 
        Parse.!!! (Parse_Spec.opt_thm_name ":" -- (Parse.term -- (@{keyword "is"} |-- Parse.term)))
 | 
|
212  | 
>> 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
 | 
213  | 
|
| 60669 | 214  | 
end;  |