| author | wenzelm | 
| Wed, 07 Aug 2024 15:11:54 +0200 | |
| changeset 80662 | ad9647592a81 | 
| parent 80634 | a90ab1ea6458 | 
| child 80673 | 5aa376b7abb8 | 
| permissions | -rw-r--r-- | 
| 47308 | 1 | (* Title: HOL/Tools/Lifting/lifting_term.ML | 
| 2 | Author: Ondrej Kuncar | |
| 3 | ||
| 4 | Proves Quotient theorem. | |
| 5 | *) | |
| 6 | ||
| 7 | signature LIFTING_TERM = | |
| 8 | sig | |
| 47379 
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
 kuncar parents: 
47350diff
changeset | 9 | exception QUOT_THM of typ * typ * Pretty.T | 
| 50227 | 10 | exception PARAM_QUOT_THM of typ * Pretty.T | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 11 | exception MERGE_TRANSFER_REL of Pretty.T | 
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 12 | exception CHECK_RTY of typ * typ | 
| 47379 
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
 kuncar parents: 
47350diff
changeset | 13 | |
| 60217 | 14 |   type 'a fold_quot_thm = { constr: typ -> thm * 'a -> thm * 'a, lift: typ -> thm * 'a -> thm * 'a, 
 | 
| 15 | comp_lift: typ -> thm * 'a -> thm * 'a } | |
| 16 | ||
| 60222 | 17 | type quotients = Lifting_Info.quotient Symtab.table | 
| 18 | ||
| 60217 | 19 | val force_qty_type: Proof.context -> typ -> thm -> thm | 
| 20 | ||
| 60222 | 21 | val prove_schematic_quot_thm: 'a fold_quot_thm -> quotients -> Proof.context -> | 
| 22 | typ * typ -> 'a -> thm * 'a | |
| 60217 | 23 | |
| 57663 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
55731diff
changeset | 24 | val instantiate_rtys: Proof.context -> typ * typ -> typ * typ | 
| 
b590fcd03a4a
store explicitly quotient types with no_code => more precise registration of code equations
 kuncar parents: 
55731diff
changeset | 25 | |
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 26 | val prove_quot_thm: Proof.context -> typ * typ -> thm | 
| 47308 | 27 | |
| 47852 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47778diff
changeset | 28 | val abs_fun: Proof.context -> typ * typ -> term | 
| 47308 | 29 | |
| 30 | val equiv_relation: Proof.context -> typ * typ -> term | |
| 50227 | 31 | |
| 50288 
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
 kuncar parents: 
50227diff
changeset | 32 | val prove_param_quot_thm: Proof.context -> typ -> thm * (typ * thm) list * Proof.context | 
| 
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
 kuncar parents: 
50227diff
changeset | 33 | |
| 
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
 kuncar parents: 
50227diff
changeset | 34 | val generate_parametrized_relator: Proof.context -> typ -> term * term list | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 35 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 36 | val merge_transfer_relations: Proof.context -> cterm -> thm | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 37 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 38 | val parametrize_transfer_rule: Proof.context -> thm -> thm | 
| 47308 | 39 | end | 
| 40 | ||
| 41 | structure Lifting_Term: LIFTING_TERM = | |
| 42 | struct | |
| 47698 | 43 | open Lifting_Util | 
| 44 | ||
| 45 | infix 0 MRSL | |
| 46 | ||
| 47379 
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
 kuncar parents: 
47350diff
changeset | 47 | exception QUOT_THM_INTERNAL of Pretty.T | 
| 
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
 kuncar parents: 
47350diff
changeset | 48 | exception QUOT_THM of typ * typ * Pretty.T | 
| 50227 | 49 | exception PARAM_QUOT_THM of typ * Pretty.T | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 50 | exception MERGE_TRANSFER_REL of Pretty.T | 
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 51 | exception CHECK_RTY of typ * typ | 
| 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 52 | |
| 60222 | 53 | type quotients = Lifting_Info.quotient Symtab.table | 
| 54 | ||
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 55 | fun match ctxt err ty_pat ty = | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 56 | Sign.typ_match (Proof_Context.theory_of ctxt) (ty_pat, ty) Vartab.empty | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 57 | handle Type.TYPE_MATCH => err ctxt ty_pat ty | 
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 58 | |
| 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 59 | fun equiv_match_err ctxt ty_pat ty = | 
| 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 60 | let | 
| 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 61 | val ty_pat_str = Syntax.string_of_typ ctxt ty_pat | 
| 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 62 | val ty_str = Syntax.string_of_typ ctxt ty | 
| 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 63 | in | 
| 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 64 | raise QUOT_THM_INTERNAL (Pretty.block | 
| 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 65 |       [Pretty.str ("The quotient type " ^ quote ty_str),
 | 
| 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 66 | Pretty.brk 1, | 
| 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 67 |        Pretty.str ("and the quotient type pattern " ^ quote ty_pat_str),
 | 
| 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 68 | Pretty.brk 1, | 
| 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 69 | Pretty.str "don't match."]) | 
| 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 70 | end | 
| 47308 | 71 | |
| 60239 | 72 | fun get_quot_data (quotients: quotients) s = | 
| 60222 | 73 | case Symtab.lookup quotients s of | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 74 | SOME qdata => qdata | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 75 | | NONE => raise QUOT_THM_INTERNAL (Pretty.block | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 76 |     [Pretty.str ("No quotient type " ^ quote s), 
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 77 | Pretty.brk 1, | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 78 | Pretty.str "found."]) | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 79 | |
| 60222 | 80 | fun get_quot_thm quotients ctxt s = | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 81 | Thm.transfer' ctxt (#quot_thm (get_quot_data quotients s)) | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 82 | |
| 60222 | 83 | fun has_pcrel_info quotients s = is_some (#pcr_info (get_quot_data quotients s)) | 
| 55455 
2cf404a469be
don't catch QOUT_THM_INTERNAL from the recursive call of parametrize_relation_conv
 kuncar parents: 
55454diff
changeset | 84 | |
| 60222 | 85 | fun get_pcrel_info quotients s = | 
| 86 | case #pcr_info (get_quot_data quotients s) of | |
| 53219 
ca237b9e4542
use only one data slot; rename print_quotmaps to print_quot_maps; tuned
 kuncar parents: 
51374diff
changeset | 87 | SOME pcr_info => pcr_info | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 88 | | NONE => raise QUOT_THM_INTERNAL (Pretty.block | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 89 |     [Pretty.str ("No parametrized correspondce relation for " ^ quote s), 
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 90 | Pretty.brk 1, | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 91 | Pretty.str "found."]) | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 92 | |
| 60222 | 93 | fun get_pcrel_def quotients ctxt s = | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 94 | Thm.transfer' ctxt (#pcrel_def (get_pcrel_info quotients s)) | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 95 | |
| 60222 | 96 | fun get_pcr_cr_eq quotients ctxt s = | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 97 | Thm.transfer' ctxt (#pcr_cr_eq (get_pcrel_info quotients s)) | 
| 47308 | 98 | |
| 99 | fun get_rel_quot_thm ctxt s = | |
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 100 | (case Lifting_Info.lookup_quot_maps ctxt s of | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 101 | SOME map_data => Thm.transfer' ctxt (#rel_quot_thm map_data) | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 102 | | NONE => raise QUOT_THM_INTERNAL (Pretty.block | 
| 47379 
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
 kuncar parents: 
47350diff
changeset | 103 |       [Pretty.str ("No relator for the type " ^ quote s), 
 | 
| 
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
 kuncar parents: 
47350diff
changeset | 104 | Pretty.brk 1, | 
| 
075d22b3a32f
detect incorrect situations; better error messages; sanity check for quot_thm in setup_lifting_infr
 kuncar parents: 
47350diff
changeset | 105 | Pretty.str "found."])) | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 106 | |
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 107 | fun get_rel_distr_rules ctxt s tm = | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 108 | (case Lifting_Info.lookup_relator_distr_data ctxt s of | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 109 | SOME rel_distr_thm => | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 110 | (case tm of | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 111 | Const (\<^const_name>\<open>POS\<close>, _) => map (Thm.transfer' ctxt) (#pos_distr_rules rel_distr_thm) | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 112 | | Const (\<^const_name>\<open>NEG\<close>, _) => map (Thm.transfer' ctxt) (#neg_distr_rules rel_distr_thm)) | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 113 | | NONE => raise QUOT_THM_INTERNAL (Pretty.block | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 114 |       [Pretty.str ("No relator distr. data for the type " ^ quote s), 
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 115 | Pretty.brk 1, | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 116 | Pretty.str "found."])) | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 117 | |
| 59848 | 118 | fun is_id_quot thm = Thm.eq_thm_prop (thm, @{thm identity_quotient})
 | 
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 119 | |
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 120 | fun zip_Tvars ctxt0 type_name rty_Tvars qty_Tvars = | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 121 | case try (get_rel_quot_thm ctxt0) type_name of | 
| 47778 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 122 | NONE => rty_Tvars ~~ qty_Tvars | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 123 | | SOME rel_quot_thm => | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 124 | let | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 125 | fun quot_term_absT quot_term = | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 126 | let | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 127 | val (_, abs, _, _) = (dest_Quotient o HOLogic.dest_Trueprop) quot_term | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 128 | in | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 129 | fastype_of abs | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 130 | end | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 131 | |
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 132 | fun equiv_univ_err ctxt ty_pat ty = | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 133 | let | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 134 | val ty_pat_str = Syntax.string_of_typ ctxt ty_pat | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 135 | val ty_str = Syntax.string_of_typ ctxt ty | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 136 | in | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 137 | raise QUOT_THM_INTERNAL (Pretty.block | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 138 |               [Pretty.str ("The type " ^ quote ty_str),
 | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 139 | Pretty.brk 1, | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 140 |                Pretty.str ("and the relator type pattern " ^ quote ty_pat_str),
 | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 141 | Pretty.brk 1, | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 142 | Pretty.str "don't unify."]) | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 143 | end | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 144 | |
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 145 | fun raw_match (TVar (v, S), T) subs = | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 146 | (case Vartab.defined subs v of | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 147 | false => Vartab.update_new (v, (S, T)) subs | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 148 | | true => subs) | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 149 | | raw_match (Type (_, Ts), Type (_, Us)) subs = | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 150 | raw_matches (Ts, Us) subs | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 151 | | raw_match _ subs = subs | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 152 | and raw_matches (T :: Ts, U :: Us) subs = raw_matches (Ts, Us) (raw_match (T, U) subs) | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 153 | | raw_matches _ subs = subs | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 154 | |
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 155 | val rty = Type (type_name, rty_Tvars) | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 156 | val qty = Type (type_name, qty_Tvars) | 
| 59582 | 157 | val rel_quot_thm_concl = (Logic.strip_imp_concl o Thm.prop_of) rel_quot_thm | 
| 47778 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 158 | val schematic_rel_absT = quot_term_absT rel_quot_thm_concl; | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 159 | val absT = rty --> qty | 
| 54946 
1c000fa0fdf5
ensure that schematic type variables are fresh in rty
 kuncar parents: 
54945diff
changeset | 160 | val schematic_absT = | 
| 
1c000fa0fdf5
ensure that schematic type variables are fresh in rty
 kuncar parents: 
54945diff
changeset | 161 | absT | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 162 | |> Logic.type_map (singleton (Variable.polymorphic ctxt0)) | 
| 54946 
1c000fa0fdf5
ensure that schematic type variables are fresh in rty
 kuncar parents: 
54945diff
changeset | 163 | |> Logic.incr_tvar (maxidx_of_typ schematic_rel_absT + 1) | 
| 
1c000fa0fdf5
ensure that schematic type variables are fresh in rty
 kuncar parents: 
54945diff
changeset | 164 | (* because absT can already contain schematic variables from rty patterns *) | 
| 47853 | 165 | val maxidx = Term.maxidx_of_typs [schematic_rel_absT, schematic_absT] | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 166 | val _ = | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 167 | Sign.typ_unify (Proof_Context.theory_of ctxt0) (schematic_rel_absT, schematic_absT) | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 168 | (Vartab.empty, maxidx) | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 169 | handle Type.TUNIFY => equiv_univ_err ctxt0 schematic_rel_absT schematic_absT | 
| 47778 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 170 | val subs = raw_match (schematic_rel_absT, absT) Vartab.empty | 
| 59582 | 171 | val rel_quot_thm_prems = (Logic.strip_imp_prems o Thm.prop_of) rel_quot_thm | 
| 47778 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 172 | in | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 173 | map (dest_funT o Envir.subst_type subs o quot_term_absT) rel_quot_thm_prems | 
| 47778 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 174 | end | 
| 
7bcdaa255a00
support Quotient map theorems with invariant parameters
 kuncar parents: 
47777diff
changeset | 175 | |
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 176 | fun gen_rty_is_TVar quotients ctxt qty = | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 177 | qty |> Tname |> get_quot_thm quotients ctxt | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 178 | |> quot_thm_rty_qty |> fst |> is_TVar | 
| 55454 | 179 | |
| 60222 | 180 | fun gen_instantiate_rtys quotients ctxt (rty, (qty as Type (qty_name, _))) = | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 181 | let | 
| 60222 | 182 | val quot_thm = get_quot_thm quotients ctxt qty_name | 
| 55454 | 183 | val (rty_pat, qty_pat) = quot_thm_rty_qty quot_thm | 
| 184 | ||
| 54945 
dcd4dcf26395
ensure that the raw type of an abstract type cannot be treated as an abstract type recursively
 kuncar parents: 
54335diff
changeset | 185 | fun inst_rty (Type (s, tys), Type (s', tys')) = | 
| 
dcd4dcf26395
ensure that the raw type of an abstract type cannot be treated as an abstract type recursively
 kuncar parents: 
54335diff
changeset | 186 | if s = s' then Type (s', map inst_rty (tys ~~ tys')) | 
| 
dcd4dcf26395
ensure that the raw type of an abstract type cannot be treated as an abstract type recursively
 kuncar parents: 
54335diff
changeset | 187 | else raise QUOT_THM_INTERNAL (Pretty.block | 
| 
dcd4dcf26395
ensure that the raw type of an abstract type cannot be treated as an abstract type recursively
 kuncar parents: 
54335diff
changeset | 188 | [Pretty.str "The type", | 
| 
dcd4dcf26395
ensure that the raw type of an abstract type cannot be treated as an abstract type recursively
 kuncar parents: 
54335diff
changeset | 189 | Pretty.brk 1, | 
| 
dcd4dcf26395
ensure that the raw type of an abstract type cannot be treated as an abstract type recursively
 kuncar parents: 
54335diff
changeset | 190 | Syntax.pretty_typ ctxt rty, | 
| 
dcd4dcf26395
ensure that the raw type of an abstract type cannot be treated as an abstract type recursively
 kuncar parents: 
54335diff
changeset | 191 | Pretty.brk 1, | 
| 
dcd4dcf26395
ensure that the raw type of an abstract type cannot be treated as an abstract type recursively
 kuncar parents: 
54335diff
changeset | 192 |            Pretty.str ("is not a raw type for the quotient type " ^ quote qty_name ^ ";"),
 | 
| 
dcd4dcf26395
ensure that the raw type of an abstract type cannot be treated as an abstract type recursively
 kuncar parents: 
54335diff
changeset | 193 | Pretty.brk 1, | 
| 
dcd4dcf26395
ensure that the raw type of an abstract type cannot be treated as an abstract type recursively
 kuncar parents: 
54335diff
changeset | 194 | Pretty.str "the correct raw type must be an instance of", | 
| 
dcd4dcf26395
ensure that the raw type of an abstract type cannot be treated as an abstract type recursively
 kuncar parents: 
54335diff
changeset | 195 | Pretty.brk 1, | 
| 
dcd4dcf26395
ensure that the raw type of an abstract type cannot be treated as an abstract type recursively
 kuncar parents: 
54335diff
changeset | 196 | Syntax.pretty_typ ctxt rty_pat]) | 
| 
dcd4dcf26395
ensure that the raw type of an abstract type cannot be treated as an abstract type recursively
 kuncar parents: 
54335diff
changeset | 197 | | inst_rty (t as Type (_, _), TFree _) = t | 
| 
dcd4dcf26395
ensure that the raw type of an abstract type cannot be treated as an abstract type recursively
 kuncar parents: 
54335diff
changeset | 198 | | inst_rty ((TVar _), rty) = rty | 
| 
dcd4dcf26395
ensure that the raw type of an abstract type cannot be treated as an abstract type recursively
 kuncar parents: 
54335diff
changeset | 199 | | inst_rty ((TFree _), rty) = rty | 
| 
dcd4dcf26395
ensure that the raw type of an abstract type cannot be treated as an abstract type recursively
 kuncar parents: 
54335diff
changeset | 200 | | inst_rty (_, _) = error "check_raw_types: we should not be here" | 
| 
dcd4dcf26395
ensure that the raw type of an abstract type cannot be treated as an abstract type recursively
 kuncar parents: 
54335diff
changeset | 201 | |
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 202 | val qtyenv = match ctxt equiv_match_err qty_pat qty | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 203 | in | 
| 55454 | 204 | (inst_rty (rty_pat, rty), Envir.subst_type qtyenv rty_pat) | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 205 | end | 
| 60222 | 206 | | gen_instantiate_rtys _ _ _ = error "gen_instantiate_rtys: not Type" | 
| 207 | ||
| 208 | fun instantiate_rtys ctxt (rty, qty) = | |
| 209 | gen_instantiate_rtys (Lifting_Info.get_quotients ctxt) ctxt (rty, qty) | |
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 210 | |
| 60217 | 211 | type 'a fold_quot_thm = { constr: typ -> thm * 'a -> thm * 'a, lift: typ -> thm * 'a -> thm * 'a, 
 | 
| 212 | comp_lift: typ -> thm * 'a -> thm * 'a } | |
| 213 | ||
| 60239 | 214 | fun prove_schematic_quot_thm (actions: 'a fold_quot_thm) quotients ctxt (rty, qty) fold_val = | 
| 55454 | 215 | let | 
| 216 | fun lifting_step (rty, qty) = | |
| 217 | let | |
| 60222 | 218 | val (rty', rtyq) = gen_instantiate_rtys quotients ctxt (rty, qty) | 
| 219 | val (rty's, rtyqs) = if gen_rty_is_TVar quotients ctxt qty then ([rty'],[rtyq]) | |
| 55454 | 220 | else (Targs rty', Targs rtyq) | 
| 60217 | 221 | val (args, fold_val) = | 
| 60222 | 222 | fold_map (prove_schematic_quot_thm actions quotients ctxt) (rty's ~~ rtyqs) fold_val | 
| 55454 | 223 | in | 
| 224 | if forall is_id_quot args | |
| 225 | then | |
| 60217 | 226 | let | 
| 60222 | 227 | val quot_thm = get_quot_thm quotients ctxt (Tname qty) | 
| 60217 | 228 | in | 
| 229 | #lift actions qty (quot_thm, fold_val) | |
| 230 | end | |
| 55454 | 231 | else | 
| 232 | let | |
| 60222 | 233 | val quot_thm = get_quot_thm quotients ctxt (Tname qty) | 
| 234 | val rel_quot_thm = if gen_rty_is_TVar quotients ctxt qty then the_single args else | |
| 55454 | 235 | args MRSL (get_rel_quot_thm ctxt (Tname rty)) | 
| 60217 | 236 |             val comp_quot_thm = [rel_quot_thm, quot_thm] MRSL @{thm Quotient_compose}
 | 
| 55454 | 237 | in | 
| 60217 | 238 | #comp_lift actions qty (comp_quot_thm, fold_val) | 
| 55454 | 239 | end | 
| 240 | end | |
| 241 | in | |
| 242 | (case (rty, qty) of | |
| 243 | (Type (s, tys), Type (s', tys')) => | |
| 244 | if s = s' | |
| 245 | then | |
| 246 | let | |
| 60217 | 247 | val (args, fold_val) = | 
| 60222 | 248 | fold_map (prove_schematic_quot_thm actions quotients ctxt) | 
| 249 | (zip_Tvars ctxt s tys tys') fold_val | |
| 55454 | 250 | in | 
| 251 | if forall is_id_quot args | |
| 252 | then | |
| 60217 | 253 |               (@{thm identity_quotient}, fold_val)
 | 
| 55454 | 254 | else | 
| 60217 | 255 | let | 
| 256 | val quot_thm = args MRSL (get_rel_quot_thm ctxt s) | |
| 257 | in | |
| 258 | #constr actions qty (quot_thm, fold_val) | |
| 259 | end | |
| 55454 | 260 | end | 
| 261 | else | |
| 262 | lifting_step (rty, qty) | |
| 263 | | (_, Type (s', tys')) => | |
| 60222 | 264 | (case try (get_quot_thm quotients ctxt) s' of | 
| 55454 | 265 | SOME quot_thm => | 
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 266 | let | 
| 55454 | 267 | val rty_pat = (fst o quot_thm_rty_qty) quot_thm | 
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 268 | in | 
| 55454 | 269 | lifting_step (rty_pat, qty) | 
| 270 | end | |
| 271 | | NONE => | |
| 272 | let | |
| 273 |               val rty_pat = Type (s', map (fn _ => TFree ("a",[])) tys')
 | |
| 274 | in | |
| 60222 | 275 | prove_schematic_quot_thm actions quotients ctxt (rty_pat, qty) fold_val | 
| 55454 | 276 | end) | 
| 60217 | 277 |       | _ => (@{thm identity_quotient}, fold_val)
 | 
| 278 | ) | |
| 55454 | 279 | end | 
| 280 | handle QUOT_THM_INTERNAL pretty_msg => raise QUOT_THM (rty, qty, pretty_msg) | |
| 47308 | 281 | |
| 59630 | 282 | fun force_qty_type ctxt qty quot_thm = | 
| 47308 | 283 | let | 
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 284 | val (_, qty_schematic) = quot_thm_rty_qty quot_thm | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 285 | val match_env = Sign.typ_match (Proof_Context.theory_of ctxt) (qty_schematic, qty) Vartab.empty | 
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60239diff
changeset | 286 | fun prep_ty (x, (S, ty)) = ((x, S), Thm.ctyp_of ctxt ty) | 
| 59630 | 287 | val ty_inst = Vartab.fold (cons o prep_ty) match_env [] | 
| 74282 | 288 | in Thm.instantiate (TVars.make ty_inst, Vars.empty) quot_thm end | 
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 289 | |
| 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 290 | fun check_rty_type ctxt rty quot_thm = | 
| 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 291 | let | 
| 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 292 | val (rty_forced, _) = quot_thm_rty_qty quot_thm | 
| 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 293 | val rty_schematic = Logic.type_map (singleton (Variable.polymorphic ctxt)) rty | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 294 | val _ = Sign.typ_match (Proof_Context.theory_of ctxt) (rty_schematic, rty_forced) Vartab.empty | 
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 295 | handle Type.TYPE_MATCH => raise CHECK_RTY (rty_schematic, rty_forced) | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 296 | in () end | 
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 297 | |
| 47852 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47778diff
changeset | 298 | (* | 
| 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47778diff
changeset | 299 | The function tries to prove that rty and qty form a quotient. | 
| 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47778diff
changeset | 300 | |
| 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47778diff
changeset | 301 | Returns: Quotient theorem; an abstract type of the theorem is exactly | 
| 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47778diff
changeset | 302 | qty, a representation type of the theorem is an instance of rty in general. | 
| 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47778diff
changeset | 303 | *) | 
| 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47778diff
changeset | 304 | |
| 60217 | 305 | |
| 306 | local | |
| 307 |   val id_actions = { constr = K I, lift = K I, comp_lift = K I }
 | |
| 308 | in | |
| 309 | fun prove_quot_thm ctxt (rty, qty) = | |
| 310 | let | |
| 60222 | 311 | val quotients = Lifting_Info.get_quotients ctxt | 
| 312 | val (schematic_quot_thm, _) = prove_schematic_quot_thm id_actions quotients ctxt (rty, qty) () | |
| 313 | val quot_thm = force_qty_type ctxt qty schematic_quot_thm | |
| 60217 | 314 | val _ = check_rty_type ctxt rty quot_thm | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 315 | in quot_thm end | 
| 60217 | 316 | end | 
| 47308 | 317 | |
| 54335 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 318 | (* | 
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 319 | Computes the composed abstraction function for rty and qty. | 
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 320 | *) | 
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 321 | |
| 47852 
0c3b8d036a5c
documentation of the Lifting package on the ML level & tuned
 kuncar parents: 
47778diff
changeset | 322 | fun abs_fun ctxt (rty, qty) = | 
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 323 | quot_thm_abs (prove_quot_thm ctxt (rty, qty)) | 
| 47350 | 324 | |
| 54335 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 325 | (* | 
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 326 | Computes the composed equivalence relation for rty and qty. | 
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 327 | *) | 
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 328 | |
| 47308 | 329 | fun equiv_relation ctxt (rty, qty) = | 
| 47504 
aa1b8a59017f
go back to the explicit compisition of quotient theorems
 kuncar parents: 
47386diff
changeset | 330 | quot_thm_rel (prove_quot_thm ctxt (rty, qty)) | 
| 47308 | 331 | |
| 50288 
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
 kuncar parents: 
50227diff
changeset | 332 | val get_fresh_Q_t = | 
| 50227 | 333 | let | 
| 69593 | 334 | val Q_t = \<^term>\<open>Trueprop (Quotient R Abs Rep T)\<close> | 
| 50227 | 335 | val frees_Q_t = Term.add_free_names Q_t [] | 
| 50288 
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
 kuncar parents: 
50227diff
changeset | 336 | val tfrees_Q_t = rev (Term.add_tfree_names Q_t []) | 
| 50227 | 337 | in | 
| 50288 
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
 kuncar parents: 
50227diff
changeset | 338 | fn ctxt => | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 339 | let | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 340 | fun rename_free_var tab (Free (name, typ)) = | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 341 | Free (the_default name (AList.lookup op= tab name), typ) | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 342 | | rename_free_var _ t = t | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 343 | |
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 344 | fun rename_free_vars tab = map_aterms (rename_free_var tab) | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 345 | |
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 346 | fun rename_free_tvars tab = | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 347 | map_types (map_type_tfree (fn (name, sort) => | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 348 | TFree (the_default name (AList.lookup op= tab name), sort))) | 
| 50288 
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
 kuncar parents: 
50227diff
changeset | 349 | |
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 350 | val (new_frees_Q_t, ctxt') = Variable.variant_fixes frees_Q_t ctxt | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 351 | val tab_frees = frees_Q_t ~~ new_frees_Q_t | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 352 | |
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 353 | val (new_tfrees_Q_t, ctxt'') = Variable.invent_types (replicate (length tfrees_Q_t) []) ctxt' | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 354 | val tab_tfrees = tfrees_Q_t ~~ (fst o split_list) new_tfrees_Q_t | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 355 | |
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 356 | val renamed_Q_t = rename_free_vars tab_frees Q_t | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 357 | val renamed_Q_t = rename_free_tvars tab_tfrees renamed_Q_t | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 358 | in | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 359 | (renamed_Q_t, ctxt'') | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 360 | end | 
| 50227 | 361 | end | 
| 362 | ||
| 54335 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 363 | (* | 
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 364 | For the given type, it proves a composed Quotient map theorem, where for each type variable | 
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 365 | extra Quotient assumption is generated. E.g., for 'a list it generates exactly | 
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 366 | the Quotient map theorem for the list type. The function generalizes this for the whole | 
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 367 | type universe. New fresh variables in the assumptions are fixed in the returned context. | 
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 368 | |
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 369 | Returns: the composed Quotient map theorem and list mapping each type variable in ty | 
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 370 | to the corresponding assumption in the returned theorem. | 
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 371 | *) | 
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 372 | |
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 373 | fun prove_param_quot_thm ctxt0 ty = | 
| 50227 | 374 | let | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 375 | fun generate (ty as Type (s, tys)) (table, ctxt) = | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 376 | if null tys then | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 377 | let | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 378 | val instantiated_id_quot_thm = | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 379 |                 Thm.instantiate' [SOME (Thm.ctyp_of ctxt ty)] [] @{thm identity_quotient}
 | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 380 | in (instantiated_id_quot_thm, (table, ctxt)) end | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 381 | else | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 382 | let val (args, table_ctxt') = fold_map generate tys (table, ctxt) | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 383 | in (args MRSL (get_rel_quot_thm ctxt s), table_ctxt') end | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 384 | | generate ty (table, ctxt) = | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 385 | if AList.defined (op =) table ty | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 386 | then (the (AList.lookup (op =) table ty), (table, ctxt)) | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 387 | else | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 388 | let | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 389 | val (Q_t, ctxt') = get_fresh_Q_t ctxt | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 390 | val Q_thm = Thm.assume (Thm.cterm_of ctxt' Q_t) | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 391 | val table' = (ty, Q_thm) :: table | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 392 | in (Q_thm, (table', ctxt')) end | 
| 50227 | 393 | |
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 394 | val (param_quot_thm, (table, ctxt1)) = generate ty ([], ctxt0) | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 395 | in (param_quot_thm, rev table, ctxt1) end | 
| 50227 | 396 | handle QUOT_THM_INTERNAL pretty_msg => raise PARAM_QUOT_THM (ty, pretty_msg) | 
| 397 | ||
| 54335 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 398 | (* | 
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 399 | It computes a parametrized relator for the given type ty. E.g., for 'a dlist: | 
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 400 | list_all2 ?R OO cr_dlist with parameters [?R]. | 
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 401 | |
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 402 | Returns: the definitional term and list of parameters (relations). | 
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 403 | *) | 
| 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 404 | |
| 50288 
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
 kuncar parents: 
50227diff
changeset | 405 | fun generate_parametrized_relator ctxt ty = | 
| 
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
 kuncar parents: 
50227diff
changeset | 406 | let | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 407 | val (quot_thm, table, ctxt') = prove_param_quot_thm ctxt ty | 
| 50288 
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
 kuncar parents: 
50227diff
changeset | 408 | val parametrized_relator = quot_thm_crel quot_thm | 
| 
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
 kuncar parents: 
50227diff
changeset | 409 | val args = map (fn (_, q_thm) => quot_thm_crel q_thm) table | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 410 | val exported_terms = Variable.exportT_terms ctxt' ctxt (parametrized_relator :: args) | 
| 50288 
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
 kuncar parents: 
50227diff
changeset | 411 | in | 
| 
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
 kuncar parents: 
50227diff
changeset | 412 | (hd exported_terms, tl exported_terms) | 
| 
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
 kuncar parents: 
50227diff
changeset | 413 | end | 
| 
986598b0efd1
parametrized correspondence relation: more robust procedure - don't ignore sorts; tuned
 kuncar parents: 
50227diff
changeset | 414 | |
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 415 | (* Parametrization *) | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 416 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 417 | local | 
| 59582 | 418 | fun get_lhs rule = (Thm.dest_fun o Thm.dest_arg o strip_imp_concl o Thm.cprop_of) rule; | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 419 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 420 |   fun no_imp _ = raise CTERM ("no implication", []);
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 421 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 422 | infix 0 else_imp | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 423 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 424 | fun (cv1 else_imp cv2) ct = | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 425 | (cv1 ct | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 426 | handle THM _ => cv2 ct | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 427 | | CTERM _ => cv2 ct | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 428 | | TERM _ => cv2 ct | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 429 | | TYPE _ => cv2 ct); | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 430 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 431 | fun first_imp cvs = fold_rev (curry op else_imp) cvs no_imp | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 432 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 433 | fun rewr_imp rule ct = | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 434 | let | 
| 59586 | 435 | val rule1 = Thm.incr_indexes (Thm.maxidx_of_cterm ct + 1) rule; | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 436 | val lhs_rule = get_lhs rule1; | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 437 | val rule2 = Thm.rename_boundvars (Thm.term_of lhs_rule) (Thm.term_of ct) rule1; | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 438 | val lhs_ct = Thm.dest_fun ct | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 439 | in | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 440 | Thm.instantiate (Thm.match (lhs_rule, lhs_ct)) rule2 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 441 |           handle Pattern.MATCH => raise CTERM ("rewr_imp", [lhs_rule, lhs_ct])
 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 442 | end | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 443 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 444 | fun rewrs_imp rules = first_imp (map rewr_imp rules) | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 445 | in | 
| 54335 
03b10317ba78
update documentation of important public ML functions in Lifting
 kuncar parents: 
53651diff
changeset | 446 | |
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 447 | fun gen_merge_transfer_relations quotients ctxt0 ctm = | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 448 | let | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 449 | val ctm = Thm.dest_arg ctm | 
| 59582 | 450 | val tm = Thm.term_of ctm | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 451 | val rel = (hd o get_args 2) tm | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 452 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 453 | fun same_constants (Const (n1,_)) (Const (n2,_)) = n1 = n2 | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 454 | | same_constants _ _ = false | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 455 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 456 | fun prove_extra_assms ctxt ctm distr_rule = | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 457 | let | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 458 | fun prove_assm assm = | 
| 70321 | 459 |             try (Goal.prove ctxt [] [] (Thm.term_of assm)) (fn {context = goal_ctxt, ...} =>
 | 
| 460 | SOLVED' (REPEAT_ALL_NEW (resolve_tac goal_ctxt (Transfer.get_transfer_raw goal_ctxt))) 1) | |
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 461 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 462 | fun is_POS_or_NEG ctm = | 
| 59582 | 463 | case (head_of o Thm.term_of o Thm.dest_arg) ctm of | 
| 69593 | 464 | Const (\<^const_name>\<open>POS\<close>, _) => true | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 465 | | Const (\<^const_name>\<open>NEG\<close>, _) => true | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 466 | | _ => false | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 467 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 468 | val inst_distr_rule = rewr_imp distr_rule ctm | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 469 | val extra_assms = filter_out is_POS_or_NEG (cprems_of inst_distr_rule) | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 470 | val proved_assms = map_interrupt prove_assm extra_assms | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 471 | in | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 472 | Option.map (curry op OF inst_distr_rule) proved_assms | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 473 | end | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 474 | handle CTERM _ => NONE | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 475 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 476 | fun cannot_merge_error_msg () = Pretty.block | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 477 | [Pretty.str "Rewriting (merging) of this term has failed:", | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 478 | Pretty.brk 1, | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 479 | Syntax.pretty_term ctxt0 rel] | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 480 | |
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 481 | in | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 482 | case get_args 2 rel of | 
| 69593 | 483 |           [Const (\<^const_name>\<open>HOL.eq\<close>, _), _] => rewrs_imp @{thms neg_eq_OO pos_eq_OO} ctm
 | 
| 484 |           | [_, Const (\<^const_name>\<open>HOL.eq\<close>, _)] => rewrs_imp @{thms neg_OO_eq pos_OO_eq} ctm
 | |
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 485 | | [_, trans_rel] => | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 486 | let | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 487 | val (rty', qty) = (relation_types o fastype_of) trans_rel | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 488 | in | 
| 55454 | 489 | if same_type_constrs (rty', qty) then | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 490 | let | 
| 80634 
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
 wenzelm parents: 
74282diff
changeset | 491 | val distr_rules = get_rel_distr_rules ctxt0 (dest_Type_name rty') (head_of tm) | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 492 | val distr_rule = get_first (prove_extra_assms ctxt0 ctm) distr_rules | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 493 | in | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 494 | case distr_rule of | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 495 | NONE => raise MERGE_TRANSFER_REL (cannot_merge_error_msg ()) | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 496 | | SOME distr_rule => | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 497 | map (gen_merge_transfer_relations quotients ctxt0) (cprems_of distr_rule) | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 498 | MRSL distr_rule | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 499 | end | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 500 | else | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 501 | let | 
| 80634 
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
 wenzelm parents: 
74282diff
changeset | 502 | val pcrel_def = get_pcrel_def quotients ctxt0 (dest_Type_name qty) | 
| 59582 | 503 | val pcrel_const = (head_of o fst o Logic.dest_equals o Thm.prop_of) pcrel_def | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 504 | in | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 505 | if same_constants pcrel_const (head_of trans_rel) then | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 506 | let | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 507 | val unfolded_ctm = Thm.rhs_of (Conv.arg1_conv (Conv.arg_conv (Conv.rewr_conv pcrel_def)) ctm) | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 508 |                       val distr_rule = rewrs_imp @{thms POS_pcr_rule NEG_pcr_rule} unfolded_ctm
 | 
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 509 | val result = (map (gen_merge_transfer_relations quotients ctxt0) | 
| 60222 | 510 | (cprems_of distr_rule)) MRSL distr_rule | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 511 | val fold_pcr_rel = Conv.rewr_conv (Thm.symmetric pcrel_def) | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 512 | in | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 513 | Conv.fconv_rule (HOLogic.Trueprop_conv (Conv.combination_conv | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 514 | (Conv.arg_conv (Conv.arg_conv fold_pcr_rel)) fold_pcr_rel)) result | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 515 | end | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 516 | else | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 517 | raise MERGE_TRANSFER_REL (Pretty.str "Non-parametric correspondence relation used.") | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 518 | end | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 519 | end | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 520 | end | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 521 | handle QUOT_THM_INTERNAL pretty_msg => raise MERGE_TRANSFER_REL pretty_msg | 
| 60222 | 522 | |
| 523 | (* | |
| 524 | ctm - of the form "[POS|NEG] (par_R OO T) t f) ?X", where par_R is a parametricity transfer | |
| 525 | relation for t and T is a transfer relation between t and f, which consists only from | |
| 526 | parametrized transfer relations (i.e., pcr_?) and equalities op=. POS or NEG encodes | |
| 527 | co-variance or contra-variance. | |
| 528 | ||
| 529 | The function merges par_R OO T using definitions of parametrized correspondence relations | |
| 530 | (e.g., (rel_S R) OO (pcr_T op=) --> pcr_T R using the definition pcr_T R = (rel_S R) OO cr_T). | |
| 531 | *) | |
| 532 | ||
| 70320 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 533 | fun merge_transfer_relations ctxt ctm = | 
| 
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
 wenzelm parents: 
69593diff
changeset | 534 | gen_merge_transfer_relations (Lifting_Info.get_quotients ctxt) ctxt ctm | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 535 | end | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 536 | |
| 60222 | 537 | fun gen_parametrize_transfer_rule quotients ctxt thm = | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 538 | let | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 539 | fun parametrize_relation_conv ctm = | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 540 | let | 
| 59582 | 541 | val (rty, qty) = (relation_types o fastype_of) (Thm.term_of ctm) | 
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 542 | in | 
| 55454 | 543 | if same_type_constrs (rty, qty) then | 
| 544 | if forall op= (Targs rty ~~ Targs qty) then | |
| 545 | Conv.all_conv ctm | |
| 546 | else | |
| 547 | all_args_conv parametrize_relation_conv ctm | |
| 548 | else | |
| 549 | if is_Type qty then | |
| 550 | let | |
| 80634 
a90ab1ea6458
tuned: more explicit dest_Type_name and dest_Type_args;
 wenzelm parents: 
74282diff
changeset | 551 | val q = dest_Type_name qty | 
| 55454 | 552 | in | 
| 553 | let | |
| 60222 | 554 | val (rty', rtyq) = gen_instantiate_rtys quotients ctxt (rty, qty) | 
| 555 | val (rty's, rtyqs) = if gen_rty_is_TVar quotients ctxt qty then ([rty'],[rtyq]) | |
| 55454 | 556 | else (Targs rty', Targs rtyq) | 
| 557 | in | |
| 558 | if forall op= (rty's ~~ rtyqs) then | |
| 559 | let | |
| 60222 | 560 | val pcr_cr_eq = (Thm.symmetric o mk_meta_eq) (get_pcr_cr_eq quotients ctxt q) | 
| 55454 | 561 | in | 
| 562 | Conv.rewr_conv pcr_cr_eq ctm | |
| 563 | end | |
| 564 | handle QUOT_THM_INTERNAL _ => Conv.all_conv ctm | |
| 565 | else | |
| 60222 | 566 | if has_pcrel_info quotients q then | 
| 55455 
2cf404a469be
don't catch QOUT_THM_INTERNAL from the recursive call of parametrize_relation_conv
 kuncar parents: 
55454diff
changeset | 567 | let | 
| 60222 | 568 | val pcrel_def = Thm.symmetric (get_pcrel_def quotients ctxt q) | 
| 55455 
2cf404a469be
don't catch QOUT_THM_INTERNAL from the recursive call of parametrize_relation_conv
 kuncar parents: 
55454diff
changeset | 569 | in | 
| 
2cf404a469be
don't catch QOUT_THM_INTERNAL from the recursive call of parametrize_relation_conv
 kuncar parents: 
55454diff
changeset | 570 | (Conv.rewr_conv pcrel_def then_conv all_args_conv parametrize_relation_conv) ctm | 
| 
2cf404a469be
don't catch QOUT_THM_INTERNAL from the recursive call of parametrize_relation_conv
 kuncar parents: 
55454diff
changeset | 571 | end | 
| 
2cf404a469be
don't catch QOUT_THM_INTERNAL from the recursive call of parametrize_relation_conv
 kuncar parents: 
55454diff
changeset | 572 | else Conv.arg1_conv (all_args_conv parametrize_relation_conv) ctm | 
| 55454 | 573 | end | 
| 574 | end | |
| 575 | else Conv.all_conv ctm | |
| 51374 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 576 | end | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 577 | in | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 578 | Conv.fconv_rule (HOLogic.Trueprop_conv (Conv.fun2_conv parametrize_relation_conv)) thm | 
| 
84d01fd733cf
lift_definition and setup_lifting generate parametric transfer rules if parametricity theorems are provided
 kuncar parents: 
50288diff
changeset | 579 | end | 
| 60222 | 580 | |
| 581 | (* | |
| 582 | It replaces cr_T by pcr_T op= in the transfer relation. For composed | |
| 583 | abstract types, it replaces T_rel R OO cr_T by pcr_T R. If the parametrized | |
| 584 | correspondce relation does not exist, the original relation is kept. | |
| 585 | ||
| 586 | thm - a transfer rule | |
| 587 | *) | |
| 588 | ||
| 589 | fun parametrize_transfer_rule ctxt thm = | |
| 590 | gen_parametrize_transfer_rule (Lifting_Info.get_quotients ctxt) ctxt thm | |
| 591 | ||
| 53651 | 592 | end |