| author | desharna | 
| Mon, 23 May 2022 10:23:33 +0200 | |
| changeset 75460 | 7c2fe41f5ee8 | 
| parent 74525 | c960bfcb91db | 
| child 77879 | dd222e2af01a | 
| permissions | -rw-r--r-- | 
| 37744 | 1 | (* Title: HOL/Tools/Quotient/quotient_tacs.ML | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 2 | Author: Cezary Kaliszyk and Christian Urban | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 3 | |
| 35788 | 4 | Tactics for solving goal arising from lifting theorems to quotient | 
| 5 | types. | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 6 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 7 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 8 | signature QUOTIENT_TACS = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 9 | sig | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 10 | val regularize_tac: Proof.context -> int -> tactic | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 11 | val injection_tac: Proof.context -> int -> tactic | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 12 | val all_injection_tac: Proof.context -> int -> tactic | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 13 | val clean_tac: Proof.context -> int -> tactic | 
| 41444 | 14 | |
| 38859 
053c69cb4a0e
quotient package: lemmas to be lifted and descended can be pre-simplified
 Christian Urban <urbanc@in.tum.de> parents: 
38719diff
changeset | 15 | val descend_procedure_tac: Proof.context -> thm list -> int -> tactic | 
| 
053c69cb4a0e
quotient package: lemmas to be lifted and descended can be pre-simplified
 Christian Urban <urbanc@in.tum.de> parents: 
38719diff
changeset | 16 | val descend_tac: Proof.context -> thm list -> int -> tactic | 
| 45782 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 17 | val partiality_descend_procedure_tac: Proof.context -> thm list -> int -> tactic | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 18 | val partiality_descend_tac: Proof.context -> thm list -> int -> tactic | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 19 | |
| 41444 | 20 | |
| 38859 
053c69cb4a0e
quotient package: lemmas to be lifted and descended can be pre-simplified
 Christian Urban <urbanc@in.tum.de> parents: 
38719diff
changeset | 21 | val lift_procedure_tac: Proof.context -> thm list -> thm -> int -> tactic | 
| 
053c69cb4a0e
quotient package: lemmas to be lifted and descended can be pre-simplified
 Christian Urban <urbanc@in.tum.de> parents: 
38719diff
changeset | 22 | val lift_tac: Proof.context -> thm list -> thm list -> int -> tactic | 
| 37593 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 Christian Urban <urbanc@in.tum.de> parents: 
37592diff
changeset | 23 | |
| 38625 
b97bd93fb9e2
allow for pre-simplification of lifted theorems
 Christian Urban <urbanc@in.tum.de> parents: 
38624diff
changeset | 24 | val lifted: Proof.context -> typ list -> thm list -> thm -> thm | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 25 | val lifted_attrib: attribute | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 26 | end; | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 27 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 28 | structure Quotient_Tacs: QUOTIENT_TACS = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 29 | struct | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 30 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 31 | (** various helper fuctions **) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 32 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 33 | (* Since HOL_basic_ss is too "big" for us, we *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 34 | (* need to set up our own minimal simpset. *) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47308diff
changeset | 35 | fun mk_minimal_simpset ctxt = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47308diff
changeset | 36 | empty_simpset ctxt | 
| 45625 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45350diff
changeset | 37 | |> Simplifier.set_subgoaler asm_simp_tac | 
| 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45350diff
changeset | 38 | |> Simplifier.set_mksimps (mksimps []) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 39 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 40 | (* composition of two theorems, used in maps *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 41 | fun OF1 thm1 thm2 = thm2 RS thm1 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 42 | |
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52230diff
changeset | 43 | fun atomize_thm ctxt thm = | 
| 41444 | 44 | let | 
| 74245 | 45 | val thm' = Thm.legacy_freezeT (Thm.forall_intr_vars thm) (* FIXME/TODO: is this proper Isar-technology? no! *) | 
| 59582 | 46 | val thm'' = Object_Logic.atomize ctxt (Thm.cprop_of thm') | 
| 41444 | 47 | in | 
| 48 |     @{thm equal_elim_rule1} OF [thm'', thm']
 | |
| 49 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 50 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 51 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 52 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 53 | (*** Regularize Tactic ***) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 54 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 55 | (** solvers for equivp and quotient assumptions **) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 56 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 57 | fun equiv_tac ctxt = | 
| 69593 | 58 | REPEAT_ALL_NEW (resolve_tac ctxt (rev (Named_Theorems.get ctxt \<^named_theorems>\<open>quot_equiv\<close>))) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 59 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47308diff
changeset | 60 | val equiv_solver = mk_solver "Equivalence goal solver" equiv_tac | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 61 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 62 | fun quotient_tac ctxt = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 63 | (REPEAT_ALL_NEW (FIRST' | 
| 60752 | 64 |     [resolve_tac ctxt @{thms identity_quotient3},
 | 
| 69593 | 65 | resolve_tac ctxt (rev (Named_Theorems.get ctxt \<^named_theorems>\<open>quot_thm\<close>))])) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 66 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47308diff
changeset | 67 | val quotient_solver = mk_solver "Quotient goal solver" quotient_tac | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 68 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 69 | fun solve_quotient_assm ctxt thm = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 70 | case Seq.pull (quotient_tac ctxt 1 thm) of | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 71 | SOME (t, _) => t | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 72 | | _ => error "Solve_quotient_assm failed. Possibly a quotient theorem is missing." | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 73 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 74 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 75 | fun get_match_inst thy pat trm = | 
| 41444 | 76 | let | 
| 58950 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 wenzelm parents: 
57960diff
changeset | 77 | val univ = Unify.matchers (Context.Theory thy) [(pat, trm)] | 
| 41444 | 78 | val SOME (env, _) = Seq.pull univ (* raises Bind, if no unifier *) (* FIXME fragile *) | 
| 79 | val tenv = Vartab.dest (Envir.term_env env) | |
| 80 | val tyenv = Vartab.dest (Envir.type_env env) | |
| 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: 
59638diff
changeset | 81 | fun prep_ty (x, (S, ty)) = ((x, S), Thm.global_ctyp_of thy ty) | 
| 
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: 
59638diff
changeset | 82 | fun prep_trm (x, (T, t)) = ((x, T), Thm.global_cterm_of thy t) | 
| 41444 | 83 | in | 
| 74282 | 84 | (TVars.make (map prep_ty tyenv), Vars.make (map prep_trm tenv)) | 
| 41444 | 85 | end | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 86 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 87 | (* Calculates the instantiations for the lemmas: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 88 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 89 | ball_reg_eqv_range and bex_reg_eqv_range | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 90 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 91 | Since the left-hand-side contains a non-pattern '?P (f ?x)' | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 92 | we rely on unification/instantiation to check whether the | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 93 | theorem applies and return NONE if it doesn't. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 94 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 95 | fun calculate_inst ctxt ball_bex_thm redex R1 R2 = | 
| 41444 | 96 | let | 
| 42361 | 97 | val thy = Proof_Context.theory_of ctxt | 
| 41444 | 98 | fun get_lhs thm = fst (Logic.dest_equals (Thm.concl_of thm)) | 
| 59638 | 99 | val ty_inst = map (SOME o Thm.ctyp_of ctxt) [domain_type (fastype_of R2)] | 
| 100 | val trm_inst = map (SOME o Thm.cterm_of ctxt) [R2, R1] | |
| 41444 | 101 | in | 
| 60801 | 102 | (case try (Thm.instantiate' ty_inst trm_inst) ball_bex_thm of | 
| 41444 | 103 | NONE => NONE | 
| 104 | | SOME thm' => | |
| 61144 | 105 | (case try (get_match_inst thy (get_lhs thm')) (Thm.term_of redex) of | 
| 41444 | 106 | NONE => NONE | 
| 43333 
2bdec7f430d3
renamed Drule.instantiate to Drule.instantiate_normalize to emphasize its meaning as opposed to plain Thm.instantiate;
 wenzelm parents: 
42361diff
changeset | 107 | | SOME inst2 => try (Drule.instantiate_normalize inst2) thm')) | 
| 41444 | 108 | end | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 109 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47308diff
changeset | 110 | fun ball_bex_range_simproc ctxt redex = | 
| 61144 | 111 | (case Thm.term_of redex of | 
| 69593 | 112 | (Const (\<^const_name>\<open>Ball\<close>, _) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ | 
| 113 | (Const (\<^const_name>\<open>rel_fun\<close>, _) $ R1 $ R2)) $ _) => | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47308diff
changeset | 114 |         calculate_inst ctxt @{thm ball_reg_eqv_range[THEN eq_reflection]} redex R1 R2
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 115 | |
| 69593 | 116 | | (Const (\<^const_name>\<open>Bex\<close>, _) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ | 
| 117 | (Const (\<^const_name>\<open>rel_fun\<close>, _) $ R1 $ R2)) $ _) => | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47308diff
changeset | 118 |         calculate_inst ctxt @{thm bex_reg_eqv_range[THEN eq_reflection]} redex R1 R2
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 119 | |
| 61144 | 120 | | _ => NONE) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 121 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 122 | (* Regularize works as follows: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 123 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 124 | 0. preliminary simplification step according to | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 125 | ball_reg_eqv bex_reg_eqv babs_reg_eqv ball_reg_eqv_range bex_reg_eqv_range | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 126 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 127 | 1. eliminating simple Ball/Bex instances (ball_reg_right bex_reg_left) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 128 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 129 | 2. monos | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 130 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 131 | 3. commutation rules for ball and bex (ball_all_comm bex_ex_comm) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 132 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 133 | 4. then rel-equalities, which need to be instantiated with 'eq_imp_rel' | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 134 | to avoid loops | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 135 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 136 | 5. then simplification like 0 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 137 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 138 | finally jump back to 1 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 139 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 140 | |
| 37493 
2377d246a631
Quotient package now uses Partial Equivalence instead place of equivalence
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36945diff
changeset | 141 | fun reflp_get ctxt = | 
| 59585 | 142 |   map_filter (fn th => if Thm.no_prems th then SOME (OF1 @{thm equivp_reflp} th) else NONE
 | 
| 69593 | 143 | handle THM _ => NONE) (rev (Named_Theorems.get ctxt \<^named_theorems>\<open>quot_equiv\<close>)) | 
| 37493 
2377d246a631
Quotient package now uses Partial Equivalence instead place of equivalence
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36945diff
changeset | 144 | |
| 67091 | 145 | val eq_imp_rel = @{lemma "equivp R \<Longrightarrow> a = b \<longrightarrow> R a b" by (simp add: equivp_reflp)}
 | 
| 37493 
2377d246a631
Quotient package now uses Partial Equivalence instead place of equivalence
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36945diff
changeset | 146 | |
| 57960 | 147 | fun eq_imp_rel_get ctxt = | 
| 69593 | 148 | map (OF1 eq_imp_rel) (rev (Named_Theorems.get ctxt \<^named_theorems>\<open>quot_equiv\<close>)) | 
| 37493 
2377d246a631
Quotient package now uses Partial Equivalence instead place of equivalence
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36945diff
changeset | 149 | |
| 61144 | 150 | val regularize_simproc = | 
| 69593 | 151 | Simplifier.make_simproc \<^context> "regularize" | 
| 152 |     {lhss = [\<^term>\<open>Ball (Respects (R1 ===> R2)) P\<close>, \<^term>\<open>Bex (Respects (R1 ===> R2)) P\<close>],
 | |
| 62913 | 153 | proc = K ball_bex_range_simproc}; | 
| 61144 | 154 | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 155 | fun regularize_tac ctxt = | 
| 41444 | 156 | let | 
| 157 | val simpset = | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47308diff
changeset | 158 | mk_minimal_simpset ctxt | 
| 41444 | 159 |       addsimps @{thms ball_reg_eqv bex_reg_eqv babs_reg_eqv babs_simp}
 | 
| 61144 | 160 | addsimprocs [regularize_simproc] | 
| 41444 | 161 | addSolver equiv_solver addSolver quotient_solver | 
| 162 | val eq_eqvs = eq_imp_rel_get ctxt | |
| 163 | in | |
| 164 | simp_tac simpset THEN' | |
| 44285 
dd203341fd2b
Quotient Package: Regularization: do not fail if no progress is made, leave the subgoal to the user. Injection: try assumptions before extensionality to avoid looping.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
43934diff
changeset | 165 | TRY o REPEAT_ALL_NEW (CHANGED o FIRST' | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58963diff
changeset | 166 |       [resolve_tac ctxt @{thms ball_reg_right bex_reg_left bex1_bexeq_reg},
 | 
| 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58963diff
changeset | 167 | resolve_tac ctxt (Inductive.get_monos ctxt), | 
| 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58963diff
changeset | 168 |        resolve_tac ctxt @{thms ball_all_comm bex_ex_comm},
 | 
| 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58963diff
changeset | 169 | resolve_tac ctxt eq_eqvs, | 
| 41444 | 170 | simp_tac simpset]) | 
| 171 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 172 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 173 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 174 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 175 | (*** Injection Tactic ***) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 176 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 177 | (* Looks for Quot_True assumptions, and in case its parameter | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 178 | is an application, it returns the function and the argument. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 179 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 180 | fun find_qt_asm asms = | 
| 41444 | 181 | let | 
| 182 | fun find_fun trm = | |
| 183 | (case trm of | |
| 69593 | 184 | (Const (\<^const_name>\<open>Trueprop\<close>, _) $ (Const (\<^const_name>\<open>Quot_True\<close>, _) $ _)) => true | 
| 41444 | 185 | | _ => false) | 
| 186 | in | |
| 187 | (case find_first find_fun asms of | |
| 188 | SOME (_ $ (_ $ (f $ a))) => SOME (f, a) | |
| 189 | | _ => NONE) | |
| 190 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 191 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 192 | fun quot_true_simple_conv ctxt fnctn ctrm = | 
| 59582 | 193 | (case Thm.term_of ctrm of | 
| 69593 | 194 | (Const (\<^const_name>\<open>Quot_True\<close>, _) $ x) => | 
| 41444 | 195 | let | 
| 196 | val fx = fnctn x; | |
| 59638 | 197 | val cx = Thm.cterm_of ctxt x; | 
| 198 | val cfx = Thm.cterm_of ctxt fx; | |
| 199 | val cxt = Thm.ctyp_of ctxt (fastype_of x); | |
| 200 | val cfxt = Thm.ctyp_of ctxt (fastype_of fx); | |
| 60801 | 201 |         val thm = Thm.instantiate' [SOME cxt, SOME cfxt] [SOME cx, SOME cfx] @{thm QT_imp}
 | 
| 41444 | 202 | in | 
| 203 | Conv.rewr_conv thm ctrm | |
| 59582 | 204 | end) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 205 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 206 | fun quot_true_conv ctxt fnctn ctrm = | 
| 59582 | 207 | (case Thm.term_of ctrm of | 
| 69593 | 208 | (Const (\<^const_name>\<open>Quot_True\<close>, _) $ _) => | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 209 | quot_true_simple_conv ctxt fnctn ctrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 210 | | _ $ _ => Conv.comb_conv (quot_true_conv ctxt fnctn) ctrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 211 | | Abs _ => Conv.abs_conv (fn (_, ctxt) => quot_true_conv ctxt fnctn) ctxt ctrm | 
| 41444 | 212 | | _ => Conv.all_conv ctrm) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 213 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 214 | fun quot_true_tac ctxt fnctn = | 
| 41444 | 215 | CONVERSION | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 216 | ((Conv.params_conv ~1 (fn ctxt => | 
| 41444 | 217 | (Conv.prems_conv ~1 (quot_true_conv ctxt fnctn)))) ctxt) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 218 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 219 | fun dest_comb (f $ a) = (f, a) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 220 | fun dest_bcomb ((_ $ l) $ r) = (l, r) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 221 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 222 | fun unlam t = | 
| 41444 | 223 | (case t of | 
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74282diff
changeset | 224 | Abs _ => snd (Term.dest_abs_global t) | 
| 41444 | 225 |   | _ => unlam (Abs("", domain_type (fastype_of t), (incr_boundvars 1 t) $ (Bound 0))))
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 226 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 227 | val bare_concl = HOLogic.dest_Trueprop o Logic.strip_assums_concl | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 228 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 229 | (* We apply apply_rsp only in case if the type needs lifting. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 230 | This is the case if the type of the data in the Quot_True | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 231 | assumption is different from the corresponding type in the goal. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 232 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 233 | val apply_rsp_tac = | 
| 59638 | 234 |   Subgoal.FOCUS (fn {concl, asms, context = ctxt,...} =>
 | 
| 41444 | 235 | let | 
| 59582 | 236 | val bare_concl = HOLogic.dest_Trueprop (Thm.term_of concl) | 
| 237 | val qt_asm = find_qt_asm (map Thm.term_of asms) | |
| 41444 | 238 | in | 
| 239 | case (bare_concl, qt_asm) of | |
| 240 | (R2 $ (f $ x) $ (g $ y), SOME (qt_fun, qt_arg)) => | |
| 241 | if fastype_of qt_fun = fastype_of f | |
| 242 | then no_tac | |
| 243 | else | |
| 244 | let | |
| 245 | val ty_x = fastype_of x | |
| 246 | val ty_b = fastype_of qt_arg | |
| 247 | val ty_f = range_type (fastype_of f) | |
| 59638 | 248 | val ty_inst = map (SOME o Thm.ctyp_of ctxt) [ty_x, ty_b, ty_f] | 
| 249 | val t_inst = map (SOME o Thm.cterm_of ctxt) [R2, f, g, x, y]; | |
| 60801 | 250 | val inst_thm = Thm.instantiate' ty_inst | 
| 47308 | 251 |                 ([NONE, NONE, NONE] @ t_inst) @{thm apply_rspQ3}
 | 
| 41444 | 252 | in | 
| 60752 | 253 | (resolve_tac ctxt [inst_thm] THEN' SOLVED' (quotient_tac ctxt)) 1 | 
| 41444 | 254 | end | 
| 255 | | _ => no_tac | |
| 256 | end) | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 257 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 258 | (* Instantiates and applies 'equals_rsp'. Since the theorem is | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 259 | complex we rely on instantiation to tell us if it applies | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 260 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 261 | fun equals_rsp_tac R ctxt = | 
| 59638 | 262 | case try (Thm.cterm_of ctxt) R of (* There can be loose bounds in R *) (* FIXME fragile *) | 
| 263 | SOME ctm => | |
| 264 | let | |
| 265 | val ty = domain_type (fastype_of R) | |
| 266 | in | |
| 60801 | 267 | case try (Thm.instantiate' [SOME (Thm.ctyp_of ctxt ty)] | 
| 59638 | 268 |             [SOME (Thm.cterm_of ctxt R)]) @{thm equals_rsp} of
 | 
| 60752 | 269 | SOME thm => resolve_tac ctxt [thm] THEN' quotient_tac ctxt | 
| 59638 | 270 | | NONE => K no_tac | 
| 271 | end | |
| 272 | | _ => K no_tac | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 273 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 274 | fun rep_abs_rsp_tac ctxt = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 275 | SUBGOAL (fn (goal, i) => | 
| 41444 | 276 | (case try bare_concl goal of | 
| 43934 
2108763f298d
Quotient Package: handle Bound variables in rep_abs_rsp_tac not only at top-level of the goal
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
43596diff
changeset | 277 | SOME (rel $ _ $ (rep $ (abs $ _))) => | 
| 
2108763f298d
Quotient Package: handle Bound variables in rep_abs_rsp_tac not only at top-level of the goal
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
43596diff
changeset | 278 | (let | 
| 40840 | 279 | val (ty_a, ty_b) = dest_funT (fastype_of abs); | 
| 59638 | 280 | val ty_inst = map (SOME o Thm.ctyp_of ctxt) [ty_a, ty_b]; | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 281 | in | 
| 59638 | 282 | case try (map (SOME o Thm.cterm_of ctxt)) [rel, abs, rep] of | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 283 | SOME t_inst => | 
| 60801 | 284 |               (case try (Thm.instantiate' ty_inst t_inst) @{thm rep_abs_rsp} of
 | 
| 60752 | 285 | SOME inst_thm => (resolve_tac ctxt [inst_thm] THEN' quotient_tac ctxt) i | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 286 | | NONE => no_tac) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 287 | | NONE => no_tac | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 288 | end | 
| 43934 
2108763f298d
Quotient Package: handle Bound variables in rep_abs_rsp_tac not only at top-level of the goal
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
43596diff
changeset | 289 | handle TERM _ => no_tac) | 
| 41444 | 290 | | _ => no_tac)) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 291 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 292 | |
| 38718 | 293 | (* Injection means to prove that the regularized theorem implies | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 294 | the abs/rep injected one. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 295 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 296 | The deterministic part: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 297 | - remove lambdas from both sides | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 298 | - prove Ball/Bex/Babs equalities using ball_rsp, bex_rsp, babs_rsp | 
| 55945 | 299 | - prove Ball/Bex relations using rel_funI | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 300 | - reflexivity of equality | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 301 | - prove equality of relations using equals_rsp | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 302 | - use user-supplied RSP theorems | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 303 | - solve 'relation of relations' goals using quot_rel_rsp | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 304 | - remove rep_abs from the right side | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 305 | (Lambdas under respects may have left us some assumptions) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 306 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 307 | Then in order: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 308 | - split applications of lifted type (apply_rsp) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 309 | - split applications of non-lifted type (cong_tac) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 310 | - apply extentionality | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 311 | - assumption | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 312 | - reflexivity of the relation | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 313 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 314 | fun injection_match_tac ctxt = SUBGOAL (fn (goal, i) => | 
| 41444 | 315 | (case bare_concl goal of | 
| 316 | (* (R1 ===> R2) (%x...) (%x...) ----> [|R1 x y|] ==> R2 (...x) (...y) *) | |
| 69593 | 317 | (Const (\<^const_name>\<open>rel_fun\<close>, _) $ _ $ _) $ (Abs _) $ (Abs _) | 
| 60752 | 318 |         => resolve_tac ctxt @{thms rel_funI} THEN' quot_true_tac ctxt unlam
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 319 | |
| 41444 | 320 | (* (op =) (Ball...) (Ball...) ----> (op =) (...) (...) *) | 
| 69593 | 321 | | (Const (\<^const_name>\<open>HOL.eq\<close>,_) $ | 
| 322 | (Const(\<^const_name>\<open>Ball\<close>,_) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ _) $ _) $ | |
| 323 | (Const(\<^const_name>\<open>Ball\<close>,_) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ _) $ _)) | |
| 60752 | 324 |         => resolve_tac ctxt @{thms ball_rsp} THEN' dresolve_tac ctxt @{thms QT_all}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 325 | |
| 41444 | 326 | (* (R1 ===> op =) (Ball...) (Ball...) ----> [|R1 x y|] ==> (Ball...x) = (Ball...y) *) | 
| 69593 | 327 | | (Const (\<^const_name>\<open>rel_fun\<close>, _) $ _ $ _) $ | 
| 328 | (Const(\<^const_name>\<open>Ball\<close>,_) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ _) $ _) $ | |
| 329 | (Const(\<^const_name>\<open>Ball\<close>,_) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ _) $ _) | |
| 60752 | 330 |         => resolve_tac ctxt @{thms rel_funI} THEN' quot_true_tac ctxt unlam
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 331 | |
| 41444 | 332 | (* (op =) (Bex...) (Bex...) ----> (op =) (...) (...) *) | 
| 69593 | 333 | | Const (\<^const_name>\<open>HOL.eq\<close>,_) $ | 
| 334 | (Const(\<^const_name>\<open>Bex\<close>,_) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ _) $ _) $ | |
| 335 | (Const(\<^const_name>\<open>Bex\<close>,_) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ _) $ _) | |
| 60752 | 336 |         => resolve_tac ctxt @{thms bex_rsp} THEN' dresolve_tac ctxt @{thms QT_ex}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 337 | |
| 41444 | 338 | (* (R1 ===> op =) (Bex...) (Bex...) ----> [|R1 x y|] ==> (Bex...x) = (Bex...y) *) | 
| 69593 | 339 | | (Const (\<^const_name>\<open>rel_fun\<close>, _) $ _ $ _) $ | 
| 340 | (Const(\<^const_name>\<open>Bex\<close>,_) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ _) $ _) $ | |
| 341 | (Const(\<^const_name>\<open>Bex\<close>,_) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ _) $ _) | |
| 60752 | 342 |         => resolve_tac ctxt @{thms rel_funI} THEN' quot_true_tac ctxt unlam
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 343 | |
| 69593 | 344 | | (Const (\<^const_name>\<open>rel_fun\<close>, _) $ _ $ _) $ | 
| 345 | (Const(\<^const_name>\<open>Bex1_rel\<close>,_) $ _) $ (Const(\<^const_name>\<open>Bex1_rel\<close>,_) $ _) | |
| 60752 | 346 |         => resolve_tac ctxt @{thms bex1_rel_rsp} THEN' quotient_tac ctxt
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 347 | |
| 41444 | 348 | | (_ $ | 
| 69593 | 349 | (Const(\<^const_name>\<open>Babs\<close>,_) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ _) $ _) $ | 
| 350 | (Const(\<^const_name>\<open>Babs\<close>,_) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ _) $ _)) | |
| 60752 | 351 |         => resolve_tac ctxt @{thms babs_rsp} THEN' quotient_tac ctxt
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 352 | |
| 69593 | 353 | | Const (\<^const_name>\<open>HOL.eq\<close>,_) $ (R $ _ $ _) $ (_ $ _ $ _) => | 
| 60752 | 354 |      (resolve_tac ctxt @{thms refl} ORELSE'
 | 
| 41444 | 355 | (equals_rsp_tac R ctxt THEN' RANGE [ | 
| 356 | quot_true_tac ctxt (fst o dest_bcomb), quot_true_tac ctxt (snd o dest_bcomb)])) | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 357 | |
| 41444 | 358 | (* reflexivity of operators arising from Cong_tac *) | 
| 69593 | 359 |   | Const (\<^const_name>\<open>HOL.eq\<close>,_) $ _ $ _ => resolve_tac ctxt @{thms refl}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 360 | |
| 41444 | 361 | (* respectfulness of constants; in particular of a simple relation *) | 
| 55945 | 362 | | _ $ (Const _) $ (Const _) (* rel_fun, list_rel, etc but not equality *) | 
| 69593 | 363 | => resolve_tac ctxt (rev (Named_Theorems.get ctxt \<^named_theorems>\<open>quot_respect\<close>)) | 
| 57960 | 364 | THEN_ALL_NEW quotient_tac ctxt | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 365 | |
| 41444 | 366 | (* R (...) (Rep (Abs ...)) ----> R (...) (...) *) | 
| 367 | (* observe map_fun *) | |
| 368 | | _ $ _ $ _ | |
| 60752 | 369 |       => (resolve_tac ctxt @{thms quot_rel_rsp} THEN_ALL_NEW quotient_tac ctxt)
 | 
| 41444 | 370 | ORELSE' rep_abs_rsp_tac ctxt | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 371 | |
| 41444 | 372 | | _ => K no_tac) i) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 373 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 374 | fun injection_step_tac ctxt rel_refl = | 
| 41444 | 375 | FIRST' [ | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 376 | injection_match_tac ctxt, | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 377 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 378 | (* R (t $ ...) (t' $ ...) ----> apply_rsp provided type of t needs lifting *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 379 | apply_rsp_tac ctxt THEN' | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 380 | RANGE [quot_true_tac ctxt (fst o dest_comb), quot_true_tac ctxt (snd o dest_comb)], | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 381 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 382 | (* (op =) (t $ ...) (t' $ ...) ----> Cong provided type of t does not need lifting *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 383 | (* merge with previous tactic *) | 
| 58956 
a816aa3ff391
proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
 wenzelm parents: 
58950diff
changeset | 384 |     Cong_Tac.cong_tac ctxt @{thm cong} THEN'
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 385 | RANGE [quot_true_tac ctxt (fst o dest_comb), quot_true_tac ctxt (snd o dest_comb)], | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 386 | |
| 44285 
dd203341fd2b
Quotient Package: Regularization: do not fail if no progress is made, leave the subgoal to the user. Injection: try assumptions before extensionality to avoid looping.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
43934diff
changeset | 387 | (* resolving with R x y assumptions *) | 
| 58963 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 wenzelm parents: 
58956diff
changeset | 388 | assume_tac ctxt, | 
| 44285 
dd203341fd2b
Quotient Package: Regularization: do not fail if no progress is made, leave the subgoal to the user. Injection: try assumptions before extensionality to avoid looping.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
43934diff
changeset | 389 | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 390 | (* (op =) (%x...) (%y...) ----> (op =) (...) (...) *) | 
| 60752 | 391 |     resolve_tac ctxt @{thms ext} THEN' quot_true_tac ctxt unlam,
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 392 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 393 | (* reflexivity of the basic relations *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 394 | (* R ... ... *) | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58963diff
changeset | 395 | resolve_tac ctxt rel_refl] | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 396 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 397 | fun injection_tac ctxt = | 
| 41444 | 398 | let | 
| 399 | val rel_refl = reflp_get ctxt | |
| 400 | in | |
| 401 | injection_step_tac ctxt rel_refl | |
| 402 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 403 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 404 | fun all_injection_tac ctxt = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 405 | REPEAT_ALL_NEW (injection_tac ctxt) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 406 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 407 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 408 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 409 | (*** Cleaning of the Theorem ***) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 410 | |
| 40602 | 411 | (* expands all map_funs, except in front of the (bound) variables listed in xs *) | 
| 412 | fun map_fun_simple_conv xs ctrm = | |
| 59582 | 413 | (case Thm.term_of ctrm of | 
| 69593 | 414 | ((Const (\<^const_name>\<open>map_fun\<close>, _) $ _ $ _) $ h $ _) => | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 415 | if member (op=) xs h | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 416 | then Conv.all_conv ctrm | 
| 40602 | 417 |         else Conv.rewr_conv @{thm map_fun_apply [THEN eq_reflection]} ctrm
 | 
| 41444 | 418 | | _ => Conv.all_conv ctrm) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 419 | |
| 40602 | 420 | fun map_fun_conv xs ctxt ctrm = | 
| 59582 | 421 | (case Thm.term_of ctrm of | 
| 41444 | 422 | _ $ _ => | 
| 423 | (Conv.comb_conv (map_fun_conv xs ctxt) then_conv | |
| 424 | map_fun_simple_conv xs) ctrm | |
| 59582 | 425 | | Abs _ => Conv.abs_conv (fn (x, ctxt) => map_fun_conv (Thm.term_of x :: xs) ctxt) ctxt ctrm | 
| 41444 | 426 | | _ => Conv.all_conv ctrm) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 427 | |
| 40602 | 428 | fun map_fun_tac ctxt = CONVERSION (map_fun_conv [] ctxt) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 429 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 430 | (* custom matching functions *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 431 | fun mk_abs u i t = | 
| 41444 | 432 | if incr_boundvars i u aconv t then Bound i | 
| 433 | else | |
| 434 | case t of | |
| 435 | t1 $ t2 => mk_abs u i t1 $ mk_abs u i t2 | |
| 436 | | Abs (s, T, t') => Abs (s, T, mk_abs u (i + 1) t') | |
| 437 | | Bound j => if i = j then error "make_inst" else t | |
| 438 | | _ => t | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 439 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 440 | fun make_inst lhs t = | 
| 41444 | 441 | let | 
| 442 |     val _ $ (Abs (_, _, (_ $ ((f as Var (_, Type ("fun", [T, _]))) $ u)))) = lhs;
 | |
| 443 | val _ $ (Abs (_, _, (_ $ g))) = t; | |
| 444 | in | |
| 445 |     (f, Abs ("x", T, mk_abs u 0 g))
 | |
| 446 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 447 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 448 | fun make_inst_id lhs t = | 
| 41444 | 449 | let | 
| 450 |     val _ $ (Abs (_, _, (f as Var (_, Type ("fun", [T, _]))) $ u)) = lhs;
 | |
| 451 | val _ $ (Abs (_, _, g)) = t; | |
| 452 | in | |
| 453 |     (f, Abs ("x", T, mk_abs u 0 g))
 | |
| 454 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 455 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 456 | (* Simplifies a redex using the 'lambda_prs' theorem. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 457 | First instantiates the types and known subterms. | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 458 | Then solves the quotient assumptions to get Rep2 and Abs1 | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 459 | Finally instantiates the function f using make_inst | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 460 | If Rep2 is an identity then the pattern is simpler and | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 461 | make_inst_id is used | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 462 | *) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 463 | fun lambda_prs_simple_conv ctxt ctrm = | 
| 59582 | 464 | (case Thm.term_of ctrm of | 
| 69593 | 465 | (Const (\<^const_name>\<open>map_fun\<close>, _) $ r1 $ a2) $ (Abs _) => | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 466 | let | 
| 40840 | 467 | val (ty_b, ty_a) = dest_funT (fastype_of r1) | 
| 468 | val (ty_c, ty_d) = dest_funT (fastype_of a2) | |
| 59638 | 469 | val tyinst = map (SOME o Thm.ctyp_of ctxt) [ty_a, ty_b, ty_c, ty_d] | 
| 470 | val tinst = [NONE, NONE, SOME (Thm.cterm_of ctxt r1), NONE, SOME (Thm.cterm_of ctxt a2)] | |
| 60801 | 471 |         val thm1 = Thm.instantiate' tyinst tinst @{thm lambda_prs[THEN eq_reflection]}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 472 | val thm2 = solve_quotient_assm ctxt (solve_quotient_assm ctxt thm1) | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52230diff
changeset | 473 |         val thm3 = rewrite_rule ctxt @{thms id_apply[THEN eq_reflection]} thm2
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 474 | val (insp, inst) = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 475 | if ty_c = ty_d | 
| 59582 | 476 | then make_inst_id (Thm.term_of (Thm.lhs_of thm3)) (Thm.term_of ctrm) | 
| 477 | else make_inst (Thm.term_of (Thm.lhs_of thm3)) (Thm.term_of ctrm) | |
| 478 | val thm4 = | |
| 74282 | 479 | Drule.instantiate_normalize | 
| 480 | (TVars.empty, Vars.make [(dest_Var insp, Thm.cterm_of ctxt inst)]) thm3 | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 481 | in | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 482 | Conv.rewr_conv thm4 ctrm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 483 | end | 
| 41444 | 484 | | _ => Conv.all_conv ctrm) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 485 | |
| 36936 
c52d1c130898
incorporated further conversions and conversionals, after some minor tuning;
 wenzelm parents: 
36850diff
changeset | 486 | fun lambda_prs_conv ctxt = Conv.top_conv lambda_prs_simple_conv ctxt | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 487 | fun lambda_prs_tac ctxt = CONVERSION (lambda_prs_conv ctxt) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 488 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 489 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 490 | (* Cleaning consists of: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 491 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 492 | 1. unfolding of ---> in front of everything, except | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 493 | bound variables (this prevents lambda_prs from | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 494 | becoming stuck) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 495 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 496 | 2. simplification with lambda_prs | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 497 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 498 | 3. simplification with: | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 499 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 500 | - Quotient_abs_rep Quotient_rel_rep | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 501 | babs_prs all_prs ex_prs ex1_prs | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 502 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 503 | - id_simps and preservation lemmas and | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 504 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 505 | - symmetric versions of the definitions | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 506 | (that is definitions of quotient constants | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 507 | are folded) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 508 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 509 | 4. test for refl | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 510 | *) | 
| 57960 | 511 | fun clean_tac ctxt = | 
| 41444 | 512 | let | 
| 57960 | 513 | val thy = Proof_Context.theory_of ctxt | 
| 45350 
257d0b179f0d
more use of global operations (see 98ec8b51af9c)
 Christian Urban <urbanc@in.tum.de> parents: 
45279diff
changeset | 514 | val defs = map (Thm.symmetric o #def) (Quotient_Info.dest_quotconsts_global thy) | 
| 69593 | 515 | val prs = rev (Named_Theorems.get ctxt \<^named_theorems>\<open>quot_preserve\<close>) | 
| 516 | val ids = rev (Named_Theorems.get ctxt \<^named_theorems>\<open>id_simps\<close>) | |
| 41444 | 517 | val thms = | 
| 47308 | 518 |       @{thms Quotient3_abs_rep Quotient3_rel_rep babs_prs all_prs ex_prs ex1_prs} @ ids @ prs @ defs
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 519 | |
| 57960 | 520 | val simpset = (mk_minimal_simpset ctxt) addsimps thms addSolver quotient_solver | 
| 41444 | 521 | in | 
| 41451 | 522 | EVERY' [ | 
| 57960 | 523 | map_fun_tac ctxt, | 
| 524 | lambda_prs_tac ctxt, | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47308diff
changeset | 525 | simp_tac simpset, | 
| 60752 | 526 | TRY o resolve_tac ctxt [refl]] | 
| 41444 | 527 | end | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 528 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 529 | |
| 38718 | 530 | (* Tactic for Generalising Free Variables in a Goal *) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 531 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 532 | fun inst_spec ctrm = | 
| 60801 | 533 |   Thm.instantiate' [SOME (Thm.ctyp_of_cterm ctrm)] [NONE, SOME ctrm] @{thm spec}
 | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 534 | |
| 60752 | 535 | fun inst_spec_tac ctxt ctrms = | 
| 536 | EVERY' (map (dresolve_tac ctxt o single o inst_spec) ctrms) | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 537 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 538 | fun all_list xs trm = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 539 | fold (fn (x, T) => fn t' => HOLogic.mk_all (x, T, t')) xs trm | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 540 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 541 | fun apply_under_Trueprop f = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 542 | HOLogic.dest_Trueprop #> f #> HOLogic.mk_Trueprop | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 543 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 544 | fun gen_frees_tac ctxt = | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 545 | SUBGOAL (fn (concl, i) => | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 546 | let | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 547 | val vrs = Term.add_frees concl [] | 
| 59638 | 548 | val cvrs = map (Thm.cterm_of ctxt o Free) vrs | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 549 | val concl' = apply_under_Trueprop (all_list vrs) concl | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 550 | val goal = Logic.mk_implies (concl', concl) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 551 | val rule = Goal.prove ctxt [] [] goal | 
| 60752 | 552 | (K (EVERY1 [inst_spec_tac ctxt (rev cvrs), assume_tac ctxt])) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 553 | in | 
| 60752 | 554 | resolve_tac ctxt [rule] i | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 555 | end) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 556 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 557 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 558 | (** The General Shape of the Lifting Procedure **) | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 559 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 560 | (* - A is the original raw theorem | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 561 | - B is the regularized theorem | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 562 | - C is the rep/abs injected version of B | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 563 | - D is the lifted theorem | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 564 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 565 | - 1st prem is the regularization step | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 566 | - 2nd prem is the rep/abs injection step | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 567 | - 3rd prem is the cleaning part | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 568 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 569 | the Quot_True premise in 2nd records the lifted theorem | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 570 | *) | 
| 45782 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 571 | val procedure_thm = | 
| 67091 | 572 |   @{lemma  "\<lbrakk>A;
 | 
| 573 | A \<longrightarrow> B; | |
| 574 | Quot_True D \<Longrightarrow> B = C; | |
| 575 | C = D\<rbrakk> \<Longrightarrow> D" | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 576 | by (simp add: Quot_True_def)} | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 577 | |
| 45782 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 578 | (* in case of partial equivalence relations, this form of the | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 579 | procedure theorem results in solvable proof obligations | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 580 | *) | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 581 | val partiality_procedure_thm = | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 582 |   @{lemma  "[|B; 
 | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 583 | Quot_True D ==> B = C; | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 584 | C = D|] ==> D" | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 585 | by (simp add: Quot_True_def)} | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 586 | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 587 | fun lift_match_error ctxt msg rtrm qtrm = | 
| 41444 | 588 | let | 
| 589 | val rtrm_str = Syntax.string_of_term ctxt rtrm | |
| 590 | val qtrm_str = Syntax.string_of_term ctxt qtrm | |
| 591 | val msg = cat_lines [enclose "[" "]" msg, "The quotient theorem", qtrm_str, | |
| 592 | "", "does not match with original theorem", rtrm_str] | |
| 593 | in | |
| 594 | error msg | |
| 595 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 596 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 597 | fun procedure_inst ctxt rtrm qtrm = | 
| 41444 | 598 | let | 
| 599 | val rtrm' = HOLogic.dest_Trueprop rtrm | |
| 600 | val qtrm' = HOLogic.dest_Trueprop qtrm | |
| 41451 | 601 | val reg_goal = Quotient_Term.regularize_trm_chk ctxt (rtrm', qtrm') | 
| 602 | handle Quotient_Term.LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm | |
| 603 | val inj_goal = Quotient_Term.inj_repabs_trm_chk ctxt (reg_goal, qtrm') | |
| 604 | handle Quotient_Term.LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm | |
| 41444 | 605 | in | 
| 60801 | 606 | Thm.instantiate' [] | 
| 59638 | 607 | [SOME (Thm.cterm_of ctxt rtrm'), | 
| 608 | SOME (Thm.cterm_of ctxt reg_goal), | |
| 41444 | 609 | NONE, | 
| 59638 | 610 | SOME (Thm.cterm_of ctxt inj_goal)] procedure_thm | 
| 41444 | 611 | end | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 612 | |
| 37593 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 Christian Urban <urbanc@in.tum.de> parents: 
37592diff
changeset | 613 | |
| 38860 
749d09f52fde
quotient package: added a list of pre-simplification rules for Ball, Bex and mem
 Christian Urban <urbanc@in.tum.de> parents: 
38859diff
changeset | 614 | (* Since we use Ball and Bex during the lifting and descending, | 
| 38862 
2795499a20bd
Quotient Package: dont unfold mem_def, use rsp and prs instead
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38860diff
changeset | 615 | we cannot deal with lemmas containing them, unless we unfold | 
| 
2795499a20bd
Quotient Package: dont unfold mem_def, use rsp and prs instead
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38860diff
changeset | 616 | them by default. *) | 
| 38860 
749d09f52fde
quotient package: added a list of pre-simplification rules for Ball, Bex and mem
 Christian Urban <urbanc@in.tum.de> parents: 
38859diff
changeset | 617 | |
| 38862 
2795499a20bd
Quotient Package: dont unfold mem_def, use rsp and prs instead
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
38860diff
changeset | 618 | val default_unfolds = @{thms Ball_def Bex_def}
 | 
| 38860 
749d09f52fde
quotient package: added a list of pre-simplification rules for Ball, Bex and mem
 Christian Urban <urbanc@in.tum.de> parents: 
38859diff
changeset | 619 | |
| 
749d09f52fde
quotient package: added a list of pre-simplification rules for Ball, Bex and mem
 Christian Urban <urbanc@in.tum.de> parents: 
38859diff
changeset | 620 | |
| 37593 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 Christian Urban <urbanc@in.tum.de> parents: 
37592diff
changeset | 621 | (** descending as tactic **) | 
| 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 Christian Urban <urbanc@in.tum.de> parents: 
37592diff
changeset | 622 | |
| 38859 
053c69cb4a0e
quotient package: lemmas to be lifted and descended can be pre-simplified
 Christian Urban <urbanc@in.tum.de> parents: 
38719diff
changeset | 623 | fun descend_procedure_tac ctxt simps = | 
| 41444 | 624 | let | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47308diff
changeset | 625 | val simpset = (mk_minimal_simpset ctxt) addsimps (simps @ default_unfolds) | 
| 41444 | 626 | in | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47308diff
changeset | 627 | full_simp_tac simpset | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52230diff
changeset | 628 | THEN' Object_Logic.full_atomize_tac ctxt | 
| 41444 | 629 | THEN' gen_frees_tac ctxt | 
| 630 | THEN' SUBGOAL (fn (goal, i) => | |
| 631 | let | |
| 45279 | 632 | val qtys = map #qtyp (Quotient_Info.dest_quotients ctxt) | 
| 41451 | 633 | val rtrm = Quotient_Term.derive_rtrm ctxt qtys goal | 
| 41444 | 634 | val rule = procedure_inst ctxt rtrm goal | 
| 635 | in | |
| 60752 | 636 | resolve_tac ctxt [rule] i | 
| 41444 | 637 | end) | 
| 638 | end | |
| 37593 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 Christian Urban <urbanc@in.tum.de> parents: 
37592diff
changeset | 639 | |
| 38859 
053c69cb4a0e
quotient package: lemmas to be lifted and descended can be pre-simplified
 Christian Urban <urbanc@in.tum.de> parents: 
38719diff
changeset | 640 | fun descend_tac ctxt simps = | 
| 41444 | 641 | let | 
| 642 | val mk_tac_raw = | |
| 643 | descend_procedure_tac ctxt simps | |
| 644 | THEN' RANGE | |
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52230diff
changeset | 645 | [Object_Logic.rulify_tac ctxt THEN' (K all_tac), | 
| 41444 | 646 | regularize_tac ctxt, | 
| 647 | all_injection_tac ctxt, | |
| 648 | clean_tac ctxt] | |
| 649 | in | |
| 650 | Goal.conjunction_tac THEN_ALL_NEW mk_tac_raw | |
| 651 | end | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 652 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 653 | |
| 45782 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 654 | (** descending for partial equivalence relations **) | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 655 | |
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 656 | fun partiality_procedure_inst ctxt rtrm qtrm = | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 657 | let | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 658 | val rtrm' = HOLogic.dest_Trueprop rtrm | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 659 | val qtrm' = HOLogic.dest_Trueprop qtrm | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 660 | val reg_goal = Quotient_Term.regularize_trm_chk ctxt (rtrm', qtrm') | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 661 | handle Quotient_Term.LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 662 | val inj_goal = Quotient_Term.inj_repabs_trm_chk ctxt (reg_goal, qtrm') | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 663 | handle Quotient_Term.LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 664 | in | 
| 60801 | 665 | Thm.instantiate' [] | 
| 59638 | 666 | [SOME (Thm.cterm_of ctxt reg_goal), | 
| 45782 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 667 | NONE, | 
| 59638 | 668 | SOME (Thm.cterm_of ctxt inj_goal)] partiality_procedure_thm | 
| 45782 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 669 | end | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 670 | |
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 671 | |
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 672 | fun partiality_descend_procedure_tac ctxt simps = | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 673 | let | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47308diff
changeset | 674 | val simpset = (mk_minimal_simpset ctxt) addsimps (simps @ default_unfolds) | 
| 45782 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 675 | in | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47308diff
changeset | 676 | full_simp_tac simpset | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52230diff
changeset | 677 | THEN' Object_Logic.full_atomize_tac ctxt | 
| 45782 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 678 | THEN' gen_frees_tac ctxt | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 679 | THEN' SUBGOAL (fn (goal, i) => | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 680 | let | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 681 | val qtys = map #qtyp (Quotient_Info.dest_quotients ctxt) | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 682 | val rtrm = Quotient_Term.derive_rtrm ctxt qtys goal | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 683 | val rule = partiality_procedure_inst ctxt rtrm goal | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 684 | in | 
| 60752 | 685 | resolve_tac ctxt [rule] i | 
| 45782 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 686 | end) | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 687 | end | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 688 | |
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 689 | fun partiality_descend_tac ctxt simps = | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 690 | let | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 691 | val mk_tac_raw = | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 692 | partiality_descend_procedure_tac ctxt simps | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 693 | THEN' RANGE | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52230diff
changeset | 694 | [Object_Logic.rulify_tac ctxt THEN' (K all_tac), | 
| 45782 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 695 | all_injection_tac ctxt, | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 696 | clean_tac ctxt] | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 697 | in | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 698 | Goal.conjunction_tac THEN_ALL_NEW mk_tac_raw | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 699 | end | 
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 700 | |
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 701 | |
| 
f82020ca3248
added a specific tactic and method that deal with partial equivalence relations
 Christian Urban <urbanc@in.tum.de> parents: 
45625diff
changeset | 702 | |
| 38625 
b97bd93fb9e2
allow for pre-simplification of lifted theorems
 Christian Urban <urbanc@in.tum.de> parents: 
38624diff
changeset | 703 | (** lifting as a tactic **) | 
| 37593 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 Christian Urban <urbanc@in.tum.de> parents: 
37592diff
changeset | 704 | |
| 38718 | 705 | |
| 37593 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 Christian Urban <urbanc@in.tum.de> parents: 
37592diff
changeset | 706 | (* the tactic leaves three subgoals to be proved *) | 
| 38859 
053c69cb4a0e
quotient package: lemmas to be lifted and descended can be pre-simplified
 Christian Urban <urbanc@in.tum.de> parents: 
38719diff
changeset | 707 | fun lift_procedure_tac ctxt simps rthm = | 
| 41444 | 708 | let | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47308diff
changeset | 709 | val simpset = (mk_minimal_simpset ctxt) addsimps (simps @ default_unfolds) | 
| 41444 | 710 | in | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47308diff
changeset | 711 | full_simp_tac simpset | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52230diff
changeset | 712 | THEN' Object_Logic.full_atomize_tac ctxt | 
| 41444 | 713 | THEN' gen_frees_tac ctxt | 
| 714 | THEN' SUBGOAL (fn (goal, i) => | |
| 715 | let | |
| 716 | (* full_atomize_tac contracts eta redexes, | |
| 717 | so we do it also in the original theorem *) | |
| 718 | val rthm' = | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47308diff
changeset | 719 | rthm |> full_simplify simpset | 
| 41444 | 720 | |> Drule.eta_contraction_rule | 
| 721 | |> Thm.forall_intr_frees | |
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52230diff
changeset | 722 | |> atomize_thm ctxt | 
| 38717 
a365f1fc5081
quotient package: deal correctly with frees in lifted theorems
 Christian Urban <urbanc@in.tum.de> parents: 
38694diff
changeset | 723 | |
| 59582 | 724 | val rule = procedure_inst ctxt (Thm.prop_of rthm') goal | 
| 41444 | 725 | in | 
| 60752 | 726 | (resolve_tac ctxt [rule] THEN' resolve_tac ctxt [rthm']) i | 
| 41444 | 727 | end) | 
| 728 | end | |
| 38625 
b97bd93fb9e2
allow for pre-simplification of lifted theorems
 Christian Urban <urbanc@in.tum.de> parents: 
38624diff
changeset | 729 | |
| 41444 | 730 | fun lift_single_tac ctxt simps rthm = | 
| 38859 
053c69cb4a0e
quotient package: lemmas to be lifted and descended can be pre-simplified
 Christian Urban <urbanc@in.tum.de> parents: 
38719diff
changeset | 731 | lift_procedure_tac ctxt simps rthm | 
| 38625 
b97bd93fb9e2
allow for pre-simplification of lifted theorems
 Christian Urban <urbanc@in.tum.de> parents: 
38624diff
changeset | 732 | THEN' RANGE | 
| 
b97bd93fb9e2
allow for pre-simplification of lifted theorems
 Christian Urban <urbanc@in.tum.de> parents: 
38624diff
changeset | 733 | [ regularize_tac ctxt, | 
| 
b97bd93fb9e2
allow for pre-simplification of lifted theorems
 Christian Urban <urbanc@in.tum.de> parents: 
38624diff
changeset | 734 | all_injection_tac ctxt, | 
| 
b97bd93fb9e2
allow for pre-simplification of lifted theorems
 Christian Urban <urbanc@in.tum.de> parents: 
38624diff
changeset | 735 | clean_tac ctxt ] | 
| 
b97bd93fb9e2
allow for pre-simplification of lifted theorems
 Christian Urban <urbanc@in.tum.de> parents: 
38624diff
changeset | 736 | |
| 38859 
053c69cb4a0e
quotient package: lemmas to be lifted and descended can be pre-simplified
 Christian Urban <urbanc@in.tum.de> parents: 
38719diff
changeset | 737 | fun lift_tac ctxt simps rthms = | 
| 41444 | 738 | Goal.conjunction_tac | 
| 38859 
053c69cb4a0e
quotient package: lemmas to be lifted and descended can be pre-simplified
 Christian Urban <urbanc@in.tum.de> parents: 
38719diff
changeset | 739 | THEN' RANGE (map (lift_single_tac ctxt simps) rthms) | 
| 38625 
b97bd93fb9e2
allow for pre-simplification of lifted theorems
 Christian Urban <urbanc@in.tum.de> parents: 
38624diff
changeset | 740 | |
| 
b97bd93fb9e2
allow for pre-simplification of lifted theorems
 Christian Urban <urbanc@in.tum.de> parents: 
38624diff
changeset | 741 | |
| 
b97bd93fb9e2
allow for pre-simplification of lifted theorems
 Christian Urban <urbanc@in.tum.de> parents: 
38624diff
changeset | 742 | (* automated lifting with pre-simplification of the theorems; | 
| 
b97bd93fb9e2
allow for pre-simplification of lifted theorems
 Christian Urban <urbanc@in.tum.de> parents: 
38624diff
changeset | 743 | for internal usage *) | 
| 
b97bd93fb9e2
allow for pre-simplification of lifted theorems
 Christian Urban <urbanc@in.tum.de> parents: 
38624diff
changeset | 744 | fun lifted ctxt qtys simps rthm = | 
| 41444 | 745 | let | 
| 746 | val ((_, [rthm']), ctxt') = Variable.import true [rthm] ctxt | |
| 59582 | 747 | val goal = Quotient_Term.derive_qtrm ctxt' qtys (Thm.prop_of rthm') | 
| 41444 | 748 | in | 
| 749 | Goal.prove ctxt' [] [] goal | |
| 750 | (K (HEADGOAL (lift_single_tac ctxt' simps rthm'))) | |
| 42361 | 751 | |> singleton (Proof_Context.export ctxt' ctxt) | 
| 41444 | 752 | end | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 753 | |
| 37593 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 Christian Urban <urbanc@in.tum.de> parents: 
37592diff
changeset | 754 | |
| 38625 
b97bd93fb9e2
allow for pre-simplification of lifted theorems
 Christian Urban <urbanc@in.tum.de> parents: 
38624diff
changeset | 755 | (* lifting as an attribute *) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 756 | |
| 61853 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 wenzelm parents: 
61144diff
changeset | 757 | val lifted_attrib = Thm.rule_attribute [] (fn context => | 
| 37593 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 Christian Urban <urbanc@in.tum.de> parents: 
37592diff
changeset | 758 | let | 
| 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 Christian Urban <urbanc@in.tum.de> parents: 
37592diff
changeset | 759 | val ctxt = Context.proof_of context | 
| 45279 | 760 | val qtys = map #qtyp (Quotient_Info.dest_quotients ctxt) | 
| 37593 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 Christian Urban <urbanc@in.tum.de> parents: 
37592diff
changeset | 761 | in | 
| 38625 
b97bd93fb9e2
allow for pre-simplification of lifted theorems
 Christian Urban <urbanc@in.tum.de> parents: 
38624diff
changeset | 762 | lifted ctxt qtys [] | 
| 37593 
2505feaf2d70
separated the lifting and descending procedures in the quotient package
 Christian Urban <urbanc@in.tum.de> parents: 
37592diff
changeset | 763 | end) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 764 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 765 | end; (* structure *) |