| author | Fabian Huch <huch@in.tum.de> | 
| Tue, 06 Aug 2024 15:00:37 +0200 | |
| changeset 80645 | a1dce0cc6c26 | 
| parent 80636 | 4041e7c8059d | 
| child 81541 | 5335b1ca6233 | 
| permissions | -rw-r--r-- | 
| 67224 | 1 | (* Title: HOL/Library/conditional_parametricity.ML | 
| 2 | Author: Jan Gilcher, Andreas Lochbihler, Dmitriy Traytel, ETH Zürich | |
| 3 | ||
| 4 | A conditional parametricity prover | |
| 5 | *) | |
| 6 | ||
| 7 | signature CONDITIONAL_PARAMETRICITY = | |
| 8 | sig | |
| 9 | exception WARNING of string | |
| 10 | type settings = | |
| 11 |     {suppress_print_theorem: bool,
 | |
| 12 | suppress_warnings: bool, | |
| 13 | warnings_as_errors: bool, | |
| 14 | use_equality_heuristic: bool} | |
| 15 | val default_settings: settings | |
| 16 | val quiet_settings: settings | |
| 17 | ||
| 18 | val parametric_constant: settings -> Attrib.binding * thm -> Proof.context -> | |
| 19 | (thm * Proof.context) | |
| 20 | val get_parametricity_theorems: Proof.context -> thm list | |
| 21 | ||
| 22 | val prove_goal: settings -> Proof.context -> thm option -> term -> thm | |
| 23 | val prove_find_goal_cond: settings -> Proof.context -> thm list -> thm option -> term -> thm | |
| 24 | ||
| 25 | val mk_goal: Proof.context -> term -> term | |
| 26 | val mk_cond_goal: Proof.context -> thm -> term * thm | |
| 27 | val mk_param_goal_from_eq_def: Proof.context -> thm -> term | |
| 28 | val step_tac: settings -> Proof.context -> thm list -> int -> tactic | |
| 29 | end | |
| 30 | ||
| 31 | structure Conditional_Parametricity: CONDITIONAL_PARAMETRICITY = | |
| 32 | struct | |
| 33 | ||
| 34 | type settings = | |
| 35 |   {suppress_print_theorem: bool,
 | |
| 36 | suppress_warnings: bool, | |
| 37 | warnings_as_errors: bool (* overrides suppress_warnings! *), | |
| 38 | use_equality_heuristic: bool}; | |
| 39 | ||
| 40 | val quiet_settings = | |
| 41 |   {suppress_print_theorem = true,
 | |
| 42 | suppress_warnings = true, | |
| 43 | warnings_as_errors = false, | |
| 44 | use_equality_heuristic = false}; | |
| 45 | ||
| 46 | val default_settings = | |
| 47 |   {suppress_print_theorem = false,
 | |
| 48 | suppress_warnings = false, | |
| 49 | warnings_as_errors = false, | |
| 50 | use_equality_heuristic = false}; | |
| 51 | ||
| 52 | (* helper functions *) | |
| 53 | ||
| 54 | fun strip_imp_prems_concl (Const("Pure.imp", _) $ A $ B) = A :: strip_imp_prems_concl B
 | |
| 55 | | strip_imp_prems_concl C = [C]; | |
| 56 | ||
| 57 | fun strip_prop_safe t = Logic.unprotect t handle TERM _ => t; | |
| 58 | ||
| 59 | fun get_class_of ctxt t = | |
| 80636 
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
 wenzelm parents: 
78465diff
changeset | 60 | Axclass.class_of_param (Proof_Context.theory_of ctxt) (dest_Const_name t); | 
| 67224 | 61 | |
| 62 | fun is_class_op ctxt t = | |
| 63 | let | |
| 64 | val t' = t |> Envir.eta_contract; | |
| 65 | in | |
| 66 | Term.is_Const t' andalso is_some (get_class_of ctxt t') | |
| 67 | end; | |
| 68 | ||
| 69 | fun apply_Var_to_bounds t = | |
| 70 | let | |
| 71 | val (t, ts) = strip_comb t; | |
| 72 | in | |
| 73 | (case t of | |
| 74 | Var (xi, _) => | |
| 75 | let | |
| 67522 | 76 | val (bounds, tail) = chop_prefix is_Bound ts; | 
| 67224 | 77 | in | 
| 78 | list_comb (Var (xi, fastype_of (betapplys (t, bounds))), map apply_Var_to_bounds tail) | |
| 79 | end | |
| 80 | | _ => list_comb (t, map apply_Var_to_bounds ts)) | |
| 81 | end; | |
| 82 | ||
| 83 | fun theorem_format_error ctxt thm = | |
| 84 | let | |
| 85 | val msg = Pretty.string_of (Pretty.chunks [(Pretty.para | |
| 86 | "Unexpected format of definition. Must be an unconditional equation."), Thm.pretty_thm ctxt thm]); | |
| 87 | in error msg end; | |
| 88 | ||
| 89 | (* Tacticals and Tactics *) | |
| 90 | ||
| 91 | exception FINISH of thm; | |
| 92 | ||
| 93 | (* Tacticals *) | |
| 94 | fun REPEAT_TRY_ELSE_DEFER tac st = | |
| 95 | let | |
| 96 | fun COMB' tac count st = ( | |
| 97 | let | |
| 98 | val n = Thm.nprems_of st; | |
| 99 | in | |
| 100 | (if n = 0 then all_tac st else | |
| 101 | (case Seq.pull ((tac THEN COMB' tac 0) st) of | |
| 102 | NONE => | |
| 103 | if count+1 = n | |
| 104 | then raise FINISH st | |
| 105 | else (defer_tac 1 THEN (COMB' tac (count+1))) st | |
| 106 | | some => Seq.make (fn () => some))) | |
| 107 | end) | |
| 108 | in COMB' tac 0 st end; | |
| 109 | ||
| 110 | (* Tactics *) | |
| 111 | (* helper tactics for printing *) | |
| 112 | fun error_tac ctxt msg st = | |
| 76051 | 113 | (error (msg ^ "\n" ^ Goal_Display.string_of_goal ctxt st); Seq.single st); | 
| 67224 | 114 | |
| 115 | fun error_tac' ctxt msg = SELECT_GOAL (error_tac ctxt msg); | |
| 116 | ||
| 117 | (* finds assumption of the form "Rel ?B Bound x Bound y", rotates it in front, | |
| 118 | applies rel_app arity times and uses ams_rl *) | |
| 119 | fun rel_app_tac ctxt t x y arity = | |
| 120 | let | |
| 121 |     val rel_app = [@{thm Rel_app}];
 | |
| 122 | val assume = [asm_rl]; | |
| 123 | fun find_and_rotate_tac t i = | |
| 124 | let | |
| 125 | fun is_correct_rule t = | |
| 126 | (case t of | |
| 69593 | 127 | Const (\<^const_name>\<open>HOL.Trueprop\<close>, _) $ (Const (\<^const_name>\<open>Transfer.Rel\<close>, _) $ | 
| 67224 | 128 | _ $ Bound x' $ Bound y') => x = x' andalso y = y' | 
| 129 | | _ => false); | |
| 130 | val idx = find_index is_correct_rule (t |> Logic.strip_assums_hyp); | |
| 131 | in | |
| 132 | if idx < 0 then no_tac else rotate_tac idx i | |
| 133 | end; | |
| 134 | fun rotate_and_dresolve_tac ctxt arity i = REPEAT_DETERM_N (arity - 1) | |
| 135 | (EVERY' [rotate_tac ~1, dresolve_tac ctxt rel_app, defer_tac] i); | |
| 136 | in | |
| 137 | SELECT_GOAL (EVERY' [find_and_rotate_tac t, forward_tac ctxt rel_app, defer_tac, | |
| 138 | rotate_and_dresolve_tac ctxt arity, rotate_tac ~1, eresolve_tac ctxt assume] 1) | |
| 139 | end; | |
| 140 | ||
| 141 | exception WARNING of string; | |
| 142 | ||
| 143 | fun transform_rules 0 thms = thms | |
| 144 | | transform_rules n thms = transform_rules (n - 1) (curry (Drule.RL o swap) | |
| 145 |       @{thms Rel_app Rel_match_app} thms);
 | |
| 146 | ||
| 147 | fun assume_equality_tac settings ctxt t arity i st = | |
| 148 | let | |
| 149 | val quiet = #suppress_warnings settings; | |
| 150 | val errors = #warnings_as_errors settings; | |
| 151 | val T = fastype_of t; | |
| 152 |     val is_eq_lemma = @{thm is_equality_Rel} |> Thm.incr_indexes ((Term.maxidx_of_term t) + 1) |>
 | |
| 153 | Drule.infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt t)]; | |
| 154 | val msg = Pretty.string_of (Pretty.chunks [Pretty.paragraph ((Pretty.text | |
| 155 | "No rule found for constant \"") @ [Syntax.pretty_term ctxt t, Pretty.str " :: " , | |
| 156 | Syntax.pretty_typ ctxt T] @ (Pretty.text "\". Using is_eq_lemma:")), Pretty.quote | |
| 157 | (Thm.pretty_thm ctxt is_eq_lemma)]); | |
| 158 | fun msg_tac st = (if errors then raise WARNING msg else if quiet then () else warning msg; | |
| 159 | Seq.single st) | |
| 160 | val tac = resolve_tac ctxt (transform_rules arity [is_eq_lemma]) i; | |
| 161 | in | |
| 162 | (if fold_atyps (K (K true)) T false then msg_tac THEN tac else tac) st | |
| 163 | end; | |
| 164 | ||
| 165 | fun mark_class_as_match_tac ctxt const const' arity = | |
| 166 | let | |
| 167 |     val rules = transform_rules arity [@{thm Rel_match_Rel} |> Thm.incr_indexes ((Int.max o
 | |
| 168 | apply2 Term.maxidx_of_term) (const, const') + 1) |> Drule.infer_instantiate' ctxt [NONE, | |
| 169 | SOME (Thm.cterm_of ctxt const), SOME (Thm.cterm_of ctxt const')]]; | |
| 170 | in resolve_tac ctxt rules end; | |
| 171 | ||
| 172 | (* transforms the parametricity theorems to fit a given arity and uses them for resolution *) | |
| 173 | fun parametricity_thm_tac settings ctxt parametricity_thms const arity = | |
| 174 | let | |
| 175 | val rules = transform_rules arity parametricity_thms; | |
| 176 | in resolve_tac ctxt rules ORELSE' assume_equality_tac settings ctxt const arity end; | |
| 177 | ||
| 178 | (* variant of parametricity_thm_tac to use matching *) | |
| 179 | fun parametricity_thm_match_tac ctxt parametricity_thms arity = | |
| 180 | let | |
| 181 | val rules = transform_rules arity parametricity_thms; | |
| 182 | in match_tac ctxt rules end; | |
| 183 | ||
| 184 | fun rel_abs_tac ctxt = resolve_tac ctxt [@{thm Rel_abs}];
 | |
| 185 | ||
| 186 | fun step_tac' settings ctxt parametricity_thms (tm, i) = | |
| 187 | (case tm |> Logic.strip_assums_concl of | |
| 69593 | 188 | Const (\<^const_name>\<open>HOL.Trueprop\<close>, _) $ (Const (rel, _) $ _ $ t $ u) => | 
| 67224 | 189 | let | 
| 190 | val (arity_of_t, arity_of_u) = apply2 (strip_comb #> snd #> length) (t, u); | |
| 191 | in | |
| 192 | (case rel of | |
| 69593 | 193 | \<^const_name>\<open>Transfer.Rel\<close> => | 
| 67224 | 194 | (case (head_of t, head_of u) of | 
| 195 | (Abs _, _) => rel_abs_tac ctxt | |
| 196 | | (_, Abs _) => rel_abs_tac ctxt | |
| 197 | | (const as (Const _), const' as (Const _)) => | |
| 198 | if #use_equality_heuristic settings andalso t aconv u | |
| 199 | then | |
| 200 | assume_equality_tac quiet_settings ctxt t 0 | |
| 201 | else if arity_of_t = arity_of_u | |
| 202 | then if is_class_op ctxt const orelse is_class_op ctxt const' | |
| 203 | then mark_class_as_match_tac ctxt const const' arity_of_t | |
| 204 | else parametricity_thm_tac settings ctxt parametricity_thms const arity_of_t | |
| 205 | else error_tac' ctxt "Malformed term. Arities of t and u don't match." | |
| 206 | | (Bound x, Bound y) => | |
| 207 | if arity_of_t = arity_of_u then if arity_of_t > 0 then rel_app_tac ctxt tm x y arity_of_t | |
| 208 | else assume_tac ctxt | |
| 209 | else error_tac' ctxt "Malformed term. Arities of t and u don't match." | |
| 210 | | _ => error_tac' ctxt | |
| 211 | "Unexpected format. Expected (Abs _, _), (_, Abs _), (Const _, Const _) or (Bound _, Bound _).") | |
| 69593 | 212 | | \<^const_name>\<open>Conditional_Parametricity.Rel_match\<close> => | 
| 67224 | 213 | parametricity_thm_match_tac ctxt parametricity_thms arity_of_t | 
| 214 | | _ => error_tac' ctxt "Unexpected format. Expected Transfer.Rel or Rel_match marker." ) i | |
| 215 | end | |
| 69593 | 216 | | Const (\<^const_name>\<open>HOL.Trueprop\<close>, _) $ (Const (\<^const_name>\<open>Transfer.is_equality\<close>, _) $ _) => | 
| 67224 | 217 | Transfer.eq_tac ctxt i | 
| 218 | | _ => error_tac' ctxt "Unexpected format. Not of form Const (HOL.Trueprop, _) $ _" i); | |
| 219 | ||
| 220 | fun step_tac settings = SUBGOAL oo step_tac' settings; | |
| 221 | ||
| 222 | fun apply_theorem_tac ctxt thm = | |
| 223 |   HEADGOAL (resolve_tac ctxt [Local_Defs.unfold ctxt @{thms Pure.prop_def} thm] THEN_ALL_NEW
 | |
| 224 | assume_tac ctxt); | |
| 225 | ||
| 226 | (* Goal Generation *) | |
| 227 | fun strip_boundvars_from_rel_match t = | |
| 228 | (case t of | |
| 69593 | 229 | (Tp as Const (\<^const_name>\<open>HOL.Trueprop\<close>, _)) $ | 
| 230 | ((Rm as Const (\<^const_name>\<open>Conditional_Parametricity.Rel_match\<close>, _)) $ R $ t $ t') => | |
| 67224 | 231 | Tp $ (Rm $ apply_Var_to_bounds R $ t $ t') | 
| 232 | | _ => t); | |
| 233 | ||
| 234 | val extract_conditions = | |
| 235 | let | |
| 236 | val filter_bounds = filter_out Term.is_open; | |
| 237 | val prem_to_conditions = | |
| 238 | map (map strip_boundvars_from_rel_match o strip_imp_prems_concl o strip_all_body); | |
| 239 | val remove_duplicates = distinct Term.aconv; | |
| 240 | in remove_duplicates o filter_bounds o flat o prem_to_conditions end; | |
| 241 | ||
| 242 | fun mk_goal ctxt t = | |
| 243 | let | |
| 244 | val ctxt = fold (Variable.declare_typ o snd) (Term.add_frees t []) ctxt; | |
| 245 | val t = singleton (Variable.polymorphic ctxt) t; | |
| 246 | val i = maxidx_of_term t + 1; | |
| 247 | fun tvar_to_tfree ((name, _), sort) = (name, sort); | |
| 248 | val tvars = Term.add_tvars t []; | |
| 249 | val new_frees = map TFree (Term.variant_frees t (map tvar_to_tfree tvars)); | |
| 250 | val u = subst_atomic_types ((map TVar tvars) ~~ new_frees) t; | |
| 251 | val T = fastype_of t; | |
| 252 | val U = fastype_of u; | |
| 69593 | 253 | val R = [T,U] ---> \<^typ>\<open>bool\<close> | 
| 67224 | 254 |     val r = Var (("R", 2 * i), R);
 | 
| 69593 | 255 | val transfer_rel = Const (\<^const_name>\<open>Transfer.Rel\<close>, [R,T,U] ---> \<^typ>\<open>bool\<close>); | 
| 67224 | 256 | in HOLogic.mk_Trueprop (transfer_rel $ r $ t $ u) end; | 
| 257 | ||
| 258 | fun mk_abs_helper T t = | |
| 259 | let | |
| 260 | val U = fastype_of t; | |
| 261 | fun mk_abs_helper' T U = | |
| 262 | if T = U then t else | |
| 263 | let | |
| 264 | val (T2, T1) = Term.dest_funT T; | |
| 265 | in | |
| 266 | Term.absdummy T2 (mk_abs_helper' T1 U) | |
| 267 | end; | |
| 268 | in mk_abs_helper' T U end; | |
| 269 | ||
| 270 | fun compare_ixs ((name, i):indexname, (name', i'):indexname) = if name < name' then LESS | |
| 271 | else if name > name' then GREATER | |
| 272 | else if i < i' then LESS | |
| 273 | else if i > i' then GREATER | |
| 274 | else EQUAL; | |
| 275 | ||
| 276 | fun mk_cond_goal ctxt thm = | |
| 277 | let | |
| 278 | val conclusion = (hd o strip_imp_prems_concl o strip_prop_safe o Thm.concl_of) thm; | |
| 279 | val conditions = (extract_conditions o Thm.prems_of) thm; | |
| 280 | val goal = Logic.list_implies (conditions, conclusion); | |
| 281 | fun compare ((ix, _), (ix', _)) = compare_ixs (ix, ix'); | |
| 282 | val goal_vars = Term.add_vars goal [] |> Ord_List.make compare; | |
| 283 | val (ixs, Ts) = split_list goal_vars; | |
| 284 | val (_, Ts') = Term.add_vars (Thm.prop_of thm) [] |> Ord_List.make compare | |
| 285 | |> Ord_List.inter compare goal_vars |> split_list; | |
| 286 | val (As, _) = Ctr_Sugar_Util.mk_Frees "A" Ts ctxt; | |
| 287 | val goal_subst = ixs ~~ As; | |
| 288 | val thm_subst = ixs ~~ (map2 mk_abs_helper Ts' As); | |
| 289 | val thm' = thm |> Drule.infer_instantiate ctxt (map (apsnd (Thm.cterm_of ctxt)) thm_subst); | |
| 290 | in (goal |> Term.subst_Vars goal_subst, thm') end; | |
| 291 | ||
| 292 | fun mk_param_goal_from_eq_def ctxt thm = | |
| 293 | let | |
| 294 | val t = | |
| 295 | (case Thm.full_prop_of thm of | |
| 69593 | 296 | (Const (\<^const_name>\<open>Pure.eq\<close>, _) $ t' $ _) => t' | 
| 67224 | 297 | | _ => theorem_format_error ctxt thm); | 
| 298 | in mk_goal ctxt t end; | |
| 299 | ||
| 300 | (* Transformations and parametricity theorems *) | |
| 301 | fun transform_class_rule ctxt thm = | |
| 302 | (case Thm.concl_of thm of | |
| 69593 | 303 | Const (\<^const_name>\<open>HOL.Trueprop\<close>, _) $ (Const (\<^const_name>\<open>Transfer.Rel\<close>, _) $ _ $ t $ u ) => | 
| 67224 | 304 | (if curry Term.aconv_untyped t u andalso is_class_op ctxt t then | 
| 305 |         thm RS @{thm Rel_Rel_match}
 | |
| 306 | else thm) | |
| 307 | | _ => thm); | |
| 308 | ||
| 309 | fun is_parametricity_theorem thm = | |
| 310 | (case Thm.concl_of thm of | |
| 69593 | 311 | Const (\<^const_name>\<open>HOL.Trueprop\<close>, _) $ (Const (rel, _) $ _ $ t $ u ) => | 
| 312 | if rel = \<^const_name>\<open>Transfer.Rel\<close> orelse | |
| 313 | rel = \<^const_name>\<open>Conditional_Parametricity.Rel_match\<close> | |
| 67224 | 314 | then curry Term.aconv_untyped t u | 
| 315 | else false | |
| 316 | | _ => false); | |
| 317 | ||
| 318 | (* Pre- and postprocessing of theorems *) | |
| 319 | fun mk_Domainp_assm (T, R) = | |
| 69593 | 320 | HOLogic.mk_eq ((Const (\<^const_name>\<open>Domainp\<close>, Term.fastype_of T --> Term.fastype_of R) $ T), R); | 
| 67224 | 321 | |
| 322 | val Domainp_lemma = | |
| 323 |   @{lemma "(!!R. Domainp T = R ==> PROP (P R)) == PROP (P (Domainp T))"
 | |
| 324 | by (rule, drule meta_spec, | |
| 325 | erule meta_mp, rule HOL.refl, simp)}; | |
| 326 | ||
| 69593 | 327 | fun fold_Domainp f (t as Const (\<^const_name>\<open>Domainp\<close>,_) $ (Var (_,_))) = f t | 
| 67224 | 328 | | fold_Domainp f (t $ u) = fold_Domainp f t #> fold_Domainp f u | 
| 329 | | fold_Domainp f (Abs (_, _, t)) = fold_Domainp f t | |
| 330 | | fold_Domainp _ _ = I; | |
| 331 | ||
| 332 | fun subst_terms tab t = | |
| 333 | let | |
| 334 | val t' = Termtab.lookup tab t | |
| 335 | in | |
| 336 | (case t' of | |
| 337 | SOME t' => t' | |
| 338 | | NONE => | |
| 339 | (case t of | |
| 340 | u $ v => (subst_terms tab u) $ (subst_terms tab v) | |
| 341 | | Abs (a, T, t) => Abs (a, T, subst_terms tab t) | |
| 342 | | t => t)) | |
| 343 | end; | |
| 344 | ||
| 345 | fun gen_abstract_domains ctxt (dest : term -> term * (term -> term)) thm = | |
| 346 | let | |
| 347 | val prop = Thm.prop_of thm | |
| 348 | val (t, mk_prop') = dest prop | |
| 349 | val Domainp_ts = rev (fold_Domainp (fn t => insert op= t) t []) | |
| 80636 
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
 wenzelm parents: 
78465diff
changeset | 350 | val Domainp_Ts = map (snd o dest_funT o dest_Const_type o fst o dest_comb) Domainp_ts | 
| 67224 | 351 | val used = Term.add_free_names t [] | 
| 352 | val rels = map (snd o dest_comb) Domainp_ts | |
| 353 | val rel_names = map (fst o fst o dest_Var) rels | |
| 354 |     val names = map (fn name => ("D" ^ name)) rel_names |> Name.variant_list used
 | |
| 355 | val frees = map Free (names ~~ Domainp_Ts) | |
| 356 | val prems = map (HOLogic.mk_Trueprop o mk_Domainp_assm) (rels ~~ frees); | |
| 357 | val t' = subst_terms (fold Termtab.update (Domainp_ts ~~ frees) Termtab.empty) t | |
| 358 | val prop1 = fold Logic.all frees (Logic.list_implies (prems, mk_prop' t')) | |
| 359 | val prop2 = Logic.list_rename_params (rev names) prop1 | |
| 360 | val cprop = Thm.cterm_of ctxt prop2 | |
| 361 | val equal_thm = Raw_Simplifier.rewrite ctxt false [Domainp_lemma] cprop | |
| 362 | fun forall_elim thm = Thm.forall_elim_vars (Thm.maxidx_of thm + 1) thm; | |
| 363 | in | |
| 364 |     forall_elim (thm COMP (equal_thm COMP @{thm equal_elim_rule2}))
 | |
| 365 | end | |
| 366 | handle TERM _ => thm; | |
| 367 | ||
| 368 | fun abstract_domains_transfer ctxt thm = | |
| 369 | let | |
| 370 | fun dest prop = | |
| 371 | let | |
| 372 | val prems = Logic.strip_imp_prems prop | |
| 373 | val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl prop) | |
| 374 | val ((rel, x), y) = apfst Term.dest_comb (Term.dest_comb concl) | |
| 375 | in | |
| 376 | (x, fn x' => | |
| 377 | Logic.list_implies (prems, HOLogic.mk_Trueprop (rel $ x' $ y))) | |
| 378 | end | |
| 379 | in | |
| 380 | gen_abstract_domains ctxt dest thm | |
| 381 | end; | |
| 382 | ||
| 383 | fun transfer_rel_conv conv = | |
| 384 | Conv.concl_conv ~1 (HOLogic.Trueprop_conv (Conv.fun2_conv (Conv.arg_conv conv))); | |
| 385 | ||
| 74545 | 386 | fun fold_relator_eqs_conv ctxt = | 
| 387 | Conv.bottom_rewrs_conv (Transfer.get_relator_eq ctxt) ctxt; | |
| 67224 | 388 | |
| 389 | fun mk_is_equality t = | |
| 69593 | 390 | Const (\<^const_name>\<open>is_equality\<close>, Term.fastype_of t --> HOLogic.boolT) $ t; | 
| 67224 | 391 | |
| 392 | val is_equality_lemma = | |
| 67399 | 393 |   @{lemma "(!!R. is_equality R ==> PROP (P R)) == PROP (P (=))"
 | 
| 67224 | 394 | by (unfold is_equality_def, rule, drule meta_spec, | 
| 395 | erule meta_mp, rule HOL.refl, simp)}; | |
| 396 | ||
| 397 | fun gen_abstract_equalities ctxt (dest : term -> term * (term -> term)) thm = | |
| 398 | let | |
| 399 | val prop = Thm.prop_of thm | |
| 400 | val (t, mk_prop') = dest prop | |
| 67399 | 401 | (* Only consider "(=)" at non-base types *) | 
| 69593 | 402 |     fun is_eq (Const (\<^const_name>\<open>HOL.eq\<close>, Type ("fun", [T, _]))) =
 | 
| 67224 | 403 | (case T of Type (_, []) => false | _ => true) | 
| 404 | | is_eq _ = false | |
| 405 | val add_eqs = Term.fold_aterms (fn t => if is_eq t then insert (op =) t else I) | |
| 406 | val eq_consts = rev (add_eqs t []) | |
| 80636 
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
 wenzelm parents: 
78465diff
changeset | 407 | val eqTs = map dest_Const_type eq_consts | 
| 67224 | 408 | val used = Term.add_free_names prop [] | 
| 409 | val names = map (K "") eqTs |> Name.variant_list used | |
| 410 | val frees = map Free (names ~~ eqTs) | |
| 411 | val prems = map (HOLogic.mk_Trueprop o mk_is_equality) frees | |
| 412 | val prop1 = mk_prop' (Term.subst_atomic (eq_consts ~~ frees) t) | |
| 413 | val prop2 = fold Logic.all frees (Logic.list_implies (prems, prop1)) | |
| 414 | val cprop = Thm.cterm_of ctxt prop2 | |
| 415 | val equal_thm = Raw_Simplifier.rewrite ctxt false [is_equality_lemma] cprop | |
| 416 | fun forall_elim thm = Thm.forall_elim_vars (Thm.maxidx_of thm + 1) thm | |
| 417 | in | |
| 418 |     forall_elim (thm COMP (equal_thm COMP @{thm equal_elim_rule2}))
 | |
| 419 | end | |
| 420 | handle TERM _ => thm; | |
| 421 | ||
| 422 | fun abstract_equalities_transfer ctxt thm = | |
| 423 | let | |
| 424 | fun dest prop = | |
| 425 | let | |
| 426 | val prems = Logic.strip_imp_prems prop | |
| 427 | val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl prop) | |
| 428 | val ((rel, x), y) = apfst Term.dest_comb (Term.dest_comb concl) | |
| 429 | in | |
| 430 | (rel, fn rel' => | |
| 431 | Logic.list_implies (prems, HOLogic.mk_Trueprop (rel' $ x $ y))) | |
| 432 | end | |
| 433 | val contracted_eq_thm = | |
| 434 | Conv.fconv_rule (transfer_rel_conv (fold_relator_eqs_conv ctxt)) thm | |
| 435 | handle CTERM _ => thm | |
| 436 | in | |
| 437 | gen_abstract_equalities ctxt dest contracted_eq_thm | |
| 438 | end; | |
| 439 | ||
| 440 | fun prep_rule ctxt = abstract_equalities_transfer ctxt #> abstract_domains_transfer ctxt; | |
| 441 | ||
| 442 | fun get_preprocess_theorems ctxt = | |
| 69593 | 443 | Named_Theorems.get ctxt \<^named_theorems>\<open>parametricity_preprocess\<close>; | 
| 67224 | 444 | |
| 445 | fun preprocess_theorem ctxt = | |
| 446 | Local_Defs.unfold0 ctxt (get_preprocess_theorems ctxt) | |
| 447 | #> transform_class_rule ctxt; | |
| 448 | ||
| 449 | fun postprocess_theorem ctxt = | |
| 450 |   Local_Defs.fold ctxt (@{thm Rel_Rel_match_eq} :: get_preprocess_theorems ctxt)
 | |
| 451 | #> prep_rule ctxt | |
| 452 |   #>  Local_Defs.unfold ctxt @{thms Rel_def};
 | |
| 453 | ||
| 454 | fun get_parametricity_theorems ctxt = | |
| 455 | let | |
| 456 | val parametricity_thm_map_filter = | |
| 457 | Option.filter (is_parametricity_theorem andf (not o curry Term.could_unify | |
| 458 |         (Thm.full_prop_of @{thm is_equality_Rel})) o Thm.full_prop_of) o preprocess_theorem ctxt;
 | |
| 459 | in | |
| 67649 | 460 | map_filter (parametricity_thm_map_filter o Thm.transfer' ctxt) | 
| 67224 | 461 | (Transfer.get_transfer_raw ctxt) | 
| 462 | end; | |
| 463 | ||
| 464 | (* Provers *) | |
| 465 | (* Tries to prove a parametricity theorem without conditions, returns the last goal_state as thm *) | |
| 466 | fun prove_find_goal_cond settings ctxt rules def_thm t = | |
| 467 | let | |
| 468 |     fun find_conditions_tac {context = ctxt, prems = _} = unfold_tac ctxt (the_list def_thm) THEN
 | |
| 469 | (REPEAT_TRY_ELSE_DEFER o HEADGOAL) (step_tac settings ctxt rules); | |
| 470 | in | |
| 471 | Goal.prove ctxt [] [] t find_conditions_tac handle FINISH st => st | |
| 472 | end; | |
| 473 | ||
| 474 | (* Simplifies and proves thm *) | |
| 475 | fun prove_cond_goal ctxt thm = | |
| 476 | let | |
| 477 | val (goal, thm') = mk_cond_goal ctxt thm; | |
| 478 | val vars = Variable.add_free_names ctxt goal []; | |
| 479 |     fun prove_conditions_tac {context = ctxt, prems = _} = apply_theorem_tac ctxt thm';
 | |
| 480 | val vars = Variable.add_free_names ctxt (Thm.prop_of thm') vars; | |
| 481 | in | |
| 482 | Goal.prove ctxt vars [] goal prove_conditions_tac | |
| 483 | end; | |
| 484 | ||
| 485 | (* Finds necessary conditions for t and proofs conditional parametricity of t under those conditions *) | |
| 486 | fun prove_goal settings ctxt def_thm t = | |
| 487 | let | |
| 488 | val parametricity_thms = get_parametricity_theorems ctxt; | |
| 489 | val found_thm = prove_find_goal_cond settings ctxt parametricity_thms def_thm t; | |
| 490 | val thm = prove_cond_goal ctxt found_thm; | |
| 491 | in | |
| 492 | postprocess_theorem ctxt thm | |
| 493 | end; | |
| 494 | ||
| 495 | (* Commands *) | |
| 496 | fun gen_parametric_constant settings prep_att prep_thm (raw_b : Attrib.binding, raw_eq) lthy = | |
| 497 | let | |
| 498 | val b = apsnd (map (prep_att lthy)) raw_b; | |
| 499 | val def_thm = (prep_thm lthy raw_eq); | |
| 500 | val eq = Ctr_Sugar_Util.mk_abs_def def_thm handle TERM _ => theorem_format_error lthy def_thm; | |
| 501 | val goal= mk_param_goal_from_eq_def lthy eq; | |
| 502 | val thm = prove_goal settings lthy (SOME eq) goal; | |
| 503 | val (res, lthy') = Local_Theory.note (b, [thm]) lthy; | |
| 504 | val _ = if #suppress_print_theorem settings then () else | |
| 78465 | 505 | Proof_Display.print_theorem (Position.thread_data ()) lthy' res; | 
| 67224 | 506 | in | 
| 507 | (the_single (snd res), lthy') | |
| 508 | end; | |
| 509 | ||
| 510 | fun parametric_constant settings = gen_parametric_constant settings (K I) (K I); | |
| 511 | ||
| 512 | val parametric_constant_cmd = snd oo gen_parametric_constant default_settings (Attrib.check_src) | |
| 513 | (singleton o Attrib.eval_thms); | |
| 514 | ||
| 515 | val _ = | |
| 69593 | 516 | Outer_Syntax.local_theory \<^command_keyword>\<open>parametric_constant\<close> "proves parametricity" | 
| 67224 | 517 | ((Parse_Spec.opt_thm_name ":" -- Parse.thm) >> parametric_constant_cmd); | 
| 518 | ||
| 67399 | 519 | end; |