src/HOL/Tools/Quotient/quotient_tacs.ML
author wenzelm
Mon, 06 Sep 2021 12:23:06 +0200
changeset 74245 282cd3aa6cc6
parent 69593 3dda49e08b9d
child 74282 c2ee8d993d6a
permissions -rw-r--r--
clarified modules;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37744
3daaf23b9ab4 tuned titles
haftmann
parents: 37593
diff changeset
     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
f1deaca15ca3 observe standard header format;
wenzelm
parents: 35625
diff changeset
     4
Tactics for solving goal arising from lifting theorems to quotient
f1deaca15ca3 observe standard header format;
wenzelm
parents: 35625
diff changeset
     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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    14
38859
053c69cb4a0e quotient package: lemmas to be lifted and descended can be pre-simplified
Christian Urban <urbanc@in.tum.de>
parents: 38719
diff 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: 38719
diff 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: 45625
diff 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: 45625
diff 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: 45625
diff changeset
    19
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    20
38859
053c69cb4a0e quotient package: lemmas to be lifted and descended can be pre-simplified
Christian Urban <urbanc@in.tum.de>
parents: 38719
diff 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: 38719
diff 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: 37592
diff changeset
    23
38625
b97bd93fb9e2 allow for pre-simplification of lifted theorems
Christian Urban <urbanc@in.tum.de>
parents: 38624
diff 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: 47308
diff changeset
    35
fun mk_minimal_simpset ctxt =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47308
diff 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: 45350
diff 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: 45350
diff 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: 52230
diff changeset
    43
fun atomize_thm ctxt thm =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    44
  let
74245
282cd3aa6cc6 clarified modules;
wenzelm
parents: 69593
diff changeset
    45
    val thm' = Thm.legacy_freezeT (Thm.forall_intr_vars thm) (* FIXME/TODO: is this proper Isar-technology? no! *)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    46
    val thm'' = Object_Logic.atomize ctxt (Thm.cprop_of thm')
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    47
  in
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    48
    @{thm equal_elim_rule1} OF [thm'', thm']
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    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
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
    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: 47308
diff 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
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
    64
    [resolve_tac ctxt @{thms identity_quotient3},
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
    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: 47308
diff 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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    76
  let
58950
d07464875dd4 optional proof context for unify operations, for the sake of proper local options;
wenzelm
parents: 57960
diff changeset
    77
    val univ = Unify.matchers (Context.Theory thy) [(pat, trm)]
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    78
    val SOME (env, _) = Seq.pull univ           (* raises Bind, if no unifier *) (* FIXME fragile *)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    79
    val tenv = Vartab.dest (Envir.term_env env)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    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: 59638
diff 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: 59638
diff changeset
    82
    fun prep_trm (x, (T, t)) = ((x, T), Thm.global_cterm_of thy t)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    83
  in
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: 59638
diff changeset
    84
    (map prep_ty tyenv, map prep_trm tenv)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    96
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41451
diff changeset
    97
    val thy = Proof_Context.theory_of ctxt
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
    98
    fun get_lhs thm = fst (Logic.dest_equals (Thm.concl_of thm))
59638
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
    99
    val ty_inst = map (SOME o Thm.ctyp_of ctxt) [domain_type (fastype_of R2)]
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   100
    val trm_inst = map (SOME o Thm.cterm_of ctxt) [R2, R1]
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   101
  in
60801
7664e0916eec tuned signature;
wenzelm
parents: 60752
diff changeset
   102
    (case try (Thm.instantiate' ty_inst trm_inst) ball_bex_thm of
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   103
      NONE => NONE
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   104
    | SOME thm' =>
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   105
        (case try (get_match_inst thy (get_lhs thm')) (Thm.term_of redex) of
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   106
          NONE => NONE
43333
2bdec7f430d3 renamed Drule.instantiate to Drule.instantiate_normalize to emphasize its meaning as opposed to plain Thm.instantiate;
wenzelm
parents: 42361
diff changeset
   107
        | SOME inst2 => try (Drule.instantiate_normalize inst2) thm'))
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   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: 47308
diff changeset
   110
fun ball_bex_range_simproc ctxt redex =
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   111
  (case Thm.term_of redex of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   112
    (Const (\<^const_name>\<open>Ball\<close>, _) $ (Const (\<^const_name>\<open>Respects\<close>, _) $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   113
      (Const (\<^const_name>\<open>rel_fun\<close>, _) $ R1 $ R2)) $ _) =>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47308
diff 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
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   116
  | (Const (\<^const_name>\<open>Bex\<close>, _) $ (Const (\<^const_name>\<open>Respects\<close>, _) $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   117
      (Const (\<^const_name>\<open>rel_fun\<close>, _) $ R1 $ R2)) $ _) =>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47308
diff 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
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   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: 36945
diff changeset
   141
fun reflp_get ctxt =
59585
wenzelm
parents: 59582
diff changeset
   142
  map_filter (fn th => if Thm.no_prems th then SOME (OF1 @{thm equivp_reflp} th) else NONE
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   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: 36945
diff changeset
   144
67091
1393c2340eec more symbols;
wenzelm
parents: 62913
diff changeset
   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: 36945
diff changeset
   146
57960
ee1ba4848896 updated to named_theorems;
wenzelm
parents: 55945
diff changeset
   147
fun eq_imp_rel_get ctxt =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   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: 36945
diff changeset
   149
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   150
val regularize_simproc =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   151
  Simplifier.make_simproc \<^context> "regularize"
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   152
    {lhss = [\<^term>\<open>Ball (Respects (R1 ===> R2)) P\<close>, \<^term>\<open>Bex (Respects (R1 ===> R2)) P\<close>],
62913
13252110a6fe eliminated unused simproc identifier;
wenzelm
parents: 61853
diff changeset
   153
     proc = K ball_bex_range_simproc};
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   154
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   155
fun regularize_tac ctxt =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   156
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   157
    val simpset =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47308
diff changeset
   158
      mk_minimal_simpset ctxt
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   159
      addsimps @{thms ball_reg_eqv bex_reg_eqv babs_reg_eqv babs_simp}
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
   160
      addsimprocs [regularize_simproc]
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   161
      addSolver equiv_solver addSolver quotient_solver
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   162
    val eq_eqvs = eq_imp_rel_get ctxt
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   163
  in
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   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: 43934
diff 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: 58963
diff 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: 58963
diff changeset
   167
       resolve_tac ctxt (Inductive.get_monos ctxt),
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58963
diff 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: 58963
diff changeset
   169
       resolve_tac ctxt eq_eqvs,
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   170
       simp_tac simpset])
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   181
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   182
    fun find_fun trm =
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   183
      (case trm of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   184
        (Const (\<^const_name>\<open>Trueprop\<close>, _) $ (Const (\<^const_name>\<open>Quot_True\<close>, _) $ _)) => true
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   185
      | _ => false)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   186
  in
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   187
     (case find_first find_fun asms of
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   188
       SOME (_ $ (_ $ (f $ a))) => SOME (f, a)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   189
     | _ => NONE)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   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
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   193
  (case Thm.term_of ctrm of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   194
    (Const (\<^const_name>\<open>Quot_True\<close>, _) $ x) =>
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   195
      let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   196
        val fx = fnctn x;
59638
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   197
        val cx = Thm.cterm_of ctxt x;
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   198
        val cfx = Thm.cterm_of ctxt fx;
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   199
        val cxt = Thm.ctyp_of ctxt (fastype_of x);
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   200
        val cfxt = Thm.ctyp_of ctxt (fastype_of fx);
60801
7664e0916eec tuned signature;
wenzelm
parents: 60752
diff changeset
   201
        val thm = Thm.instantiate' [SOME cxt, SOME cfxt] [SOME cx, SOME cfx] @{thm QT_imp}
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   202
      in
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   203
        Conv.rewr_conv thm ctrm
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   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
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   207
  (case Thm.term_of ctrm of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   223
  (case t of
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   224
    Abs a => snd (Term.dest_abs a)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   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
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   234
  Subgoal.FOCUS (fn {concl, asms, context = ctxt,...} =>
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   235
    let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   236
      val bare_concl = HOLogic.dest_Trueprop (Thm.term_of concl)
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   237
      val qt_asm = find_qt_asm (map Thm.term_of asms)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   238
    in
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   239
      case (bare_concl, qt_asm) of
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   240
        (R2 $ (f $ x) $ (g $ y), SOME (qt_fun, qt_arg)) =>
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   241
          if fastype_of qt_fun = fastype_of f
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   242
          then no_tac
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   243
          else
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   244
            let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   245
              val ty_x = fastype_of x
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   246
              val ty_b = fastype_of qt_arg
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   247
              val ty_f = range_type (fastype_of f)
59638
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   248
              val ty_inst = map (SOME o Thm.ctyp_of ctxt) [ty_x, ty_b, ty_f]
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   249
              val t_inst = map (SOME o Thm.cterm_of ctxt) [R2, f, g, x, y];
60801
7664e0916eec tuned signature;
wenzelm
parents: 60752
diff changeset
   250
              val inst_thm = Thm.instantiate' ty_inst
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 46468
diff changeset
   251
                ([NONE, NONE, NONE] @ t_inst) @{thm apply_rspQ3}
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   252
            in
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   253
              (resolve_tac ctxt [inst_thm] THEN' SOLVED' (quotient_tac ctxt)) 1
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   254
            end
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   255
      | _ => no_tac
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   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
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   262
  case try (Thm.cterm_of ctxt) R of (* There can be loose bounds in R *)  (* FIXME fragile *)
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   263
    SOME ctm =>
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   264
      let
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   265
        val ty = domain_type (fastype_of R)
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   266
      in
60801
7664e0916eec tuned signature;
wenzelm
parents: 60752
diff changeset
   267
        case try (Thm.instantiate' [SOME (Thm.ctyp_of ctxt ty)]
59638
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   268
            [SOME (Thm.cterm_of ctxt R)]) @{thm equals_rsp} of
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   269
          SOME thm => resolve_tac ctxt [thm] THEN' quotient_tac ctxt
59638
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   270
        | NONE => K no_tac
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   271
      end
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   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: 43596
diff 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: 43596
diff changeset
   278
        (let
40840
2f97215e79bf just one Term.dest_funT;
wenzelm
parents: 40602
diff changeset
   279
          val (ty_a, ty_b) = dest_funT (fastype_of abs);
59638
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   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
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   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
7664e0916eec tuned signature;
wenzelm
parents: 60752
diff changeset
   284
              (case try (Thm.instantiate' ty_inst t_inst) @{thm rep_abs_rsp} of
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   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: 43596
diff changeset
   289
        handle TERM _ => no_tac)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   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
c7cbbb18eabe tuned code
Christian Urban <urbanc@in.tum.de>
parents: 38717
diff changeset
   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
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 54742
diff changeset
   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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   315
  (case bare_concl goal of
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   316
      (* (R1 ===> R2) (%x...) (%x...) ----> [|R1 x y|] ==> R2 (...x) (...y) *)
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   317
    (Const (\<^const_name>\<open>rel_fun\<close>, _) $ _ $ _) $ (Abs _) $ (Abs _)
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   320
      (* (op =) (Ball...) (Ball...) ----> (op =) (...) (...) *)
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   321
  | (Const (\<^const_name>\<open>HOL.eq\<close>,_) $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   322
      (Const(\<^const_name>\<open>Ball\<close>,_) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ _) $ _) $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   323
      (Const(\<^const_name>\<open>Ball\<close>,_) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ _) $ _))
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   326
      (* (R1 ===> op =) (Ball...) (Ball...) ----> [|R1 x y|] ==> (Ball...x) = (Ball...y) *)
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   327
  | (Const (\<^const_name>\<open>rel_fun\<close>, _) $ _ $ _) $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   328
      (Const(\<^const_name>\<open>Ball\<close>,_) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ _) $ _) $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   329
      (Const(\<^const_name>\<open>Ball\<close>,_) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ _) $ _)
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   332
      (* (op =) (Bex...) (Bex...) ----> (op =) (...) (...) *)
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   333
  | Const (\<^const_name>\<open>HOL.eq\<close>,_) $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   334
      (Const(\<^const_name>\<open>Bex\<close>,_) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ _) $ _) $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   335
      (Const(\<^const_name>\<open>Bex\<close>,_) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ _) $ _)
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   338
      (* (R1 ===> op =) (Bex...) (Bex...) ----> [|R1 x y|] ==> (Bex...x) = (Bex...y) *)
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   339
  | (Const (\<^const_name>\<open>rel_fun\<close>, _) $ _ $ _) $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   340
      (Const(\<^const_name>\<open>Bex\<close>,_) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ _) $ _) $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   341
      (Const(\<^const_name>\<open>Bex\<close>,_) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ _) $ _)
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   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
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   344
  | (Const (\<^const_name>\<open>rel_fun\<close>, _) $ _ $ _) $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   345
      (Const(\<^const_name>\<open>Bex1_rel\<close>,_) $ _) $ (Const(\<^const_name>\<open>Bex1_rel\<close>,_) $ _)
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   348
  | (_ $
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   349
      (Const(\<^const_name>\<open>Babs\<close>,_) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ _) $ _) $
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   350
      (Const(\<^const_name>\<open>Babs\<close>,_) $ (Const (\<^const_name>\<open>Respects\<close>, _) $ _) $ _))
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   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
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   353
  | Const (\<^const_name>\<open>HOL.eq\<close>,_) $ (R $ _ $ _) $ (_ $ _ $ _) =>
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   354
     (resolve_tac ctxt @{thms refl} ORELSE'
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   355
      (equals_rsp_tac R ctxt THEN' RANGE [
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   358
      (* reflexivity of operators arising from Cong_tac *)
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   361
     (* respectfulness of constants; in particular of a simple relation *)
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 54742
diff changeset
   362
  | _ $ (Const _) $ (Const _)  (* rel_fun, list_rel, etc but not equality *)
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   363
      => resolve_tac ctxt (rev (Named_Theorems.get ctxt \<^named_theorems>\<open>quot_respect\<close>))
57960
ee1ba4848896 updated to named_theorems;
wenzelm
parents: 55945
diff changeset
   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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   366
      (* R (...) (Rep (Abs ...)) ----> R (...) (...) *)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   367
      (* observe map_fun *)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   368
  | _ $ _ $ _
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   369
      => (resolve_tac ctxt @{thms quot_rel_rsp} THEN_ALL_NEW quotient_tac ctxt)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   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: 58950
diff 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: 43934
diff changeset
   387
    (* resolving with R x y assumptions *)
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58956
diff 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: 43934
diff 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
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   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: 58963
diff 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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   398
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   399
    val rel_refl = reflp_get ctxt
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   400
  in
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   401
    injection_step_tac ctxt rel_refl
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   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
91e583511113 map_fun combinator in theory Fun
haftmann
parents: 40469
diff changeset
   411
(* expands all map_funs, except in front of the (bound) variables listed in xs *)
91e583511113 map_fun combinator in theory Fun
haftmann
parents: 40469
diff changeset
   412
fun map_fun_simple_conv xs ctrm =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   413
  (case Thm.term_of ctrm of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   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
91e583511113 map_fun combinator in theory Fun
haftmann
parents: 40469
diff changeset
   417
        else Conv.rewr_conv @{thm map_fun_apply [THEN eq_reflection]} ctrm
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   418
  | _ => Conv.all_conv ctrm)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   419
40602
91e583511113 map_fun combinator in theory Fun
haftmann
parents: 40469
diff changeset
   420
fun map_fun_conv xs ctxt ctrm =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   421
  (case Thm.term_of ctrm of
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   422
    _ $ _ =>
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   423
      (Conv.comb_conv (map_fun_conv xs ctxt) then_conv
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   424
        map_fun_simple_conv xs) ctrm
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   425
  | Abs _ => Conv.abs_conv (fn (x, ctxt) => map_fun_conv (Thm.term_of x :: xs) ctxt) ctxt ctrm
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   426
  | _ => Conv.all_conv ctrm)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   427
40602
91e583511113 map_fun combinator in theory Fun
haftmann
parents: 40469
diff changeset
   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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   432
  if incr_boundvars i u aconv t then Bound i
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   433
  else
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   434
    case t of
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   435
      t1 $ t2 => mk_abs u i t1 $ mk_abs u i t2
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   436
    | Abs (s, T, t') => Abs (s, T, mk_abs u (i + 1) t')
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   437
    | Bound j => if i = j then error "make_inst" else t
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   441
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   442
    val _ $ (Abs (_, _, (_ $ ((f as Var (_, Type ("fun", [T, _]))) $ u)))) = lhs;
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   443
    val _ $ (Abs (_, _, (_ $ g))) = t;
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   444
  in
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   445
    (f, Abs ("x", T, mk_abs u 0 g))
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   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
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   449
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   450
    val _ $ (Abs (_, _, (f as Var (_, Type ("fun", [T, _]))) $ u)) = lhs;
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   451
    val _ $ (Abs (_, _, g)) = t;
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   452
  in
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   453
    (f, Abs ("x", T, mk_abs u 0 g))
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   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
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   464
  (case Thm.term_of ctrm of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   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
2f97215e79bf just one Term.dest_funT;
wenzelm
parents: 40602
diff changeset
   467
        val (ty_b, ty_a) = dest_funT (fastype_of r1)
2f97215e79bf just one Term.dest_funT;
wenzelm
parents: 40602
diff changeset
   468
        val (ty_c, ty_d) = dest_funT (fastype_of a2)
59638
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   469
        val tyinst = map (SOME o Thm.ctyp_of ctxt) [ty_a, ty_b, ty_c, ty_d]
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   470
        val tinst = [NONE, NONE, SOME (Thm.cterm_of ctxt r1), NONE, SOME (Thm.cterm_of ctxt a2)]
60801
7664e0916eec tuned signature;
wenzelm
parents: 60752
diff changeset
   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: 52230
diff 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
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   476
          then make_inst_id (Thm.term_of (Thm.lhs_of thm3)) (Thm.term_of ctrm)
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   477
          else make_inst (Thm.term_of (Thm.lhs_of thm3)) (Thm.term_of ctrm)
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   478
        val thm4 =
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: 59638
diff changeset
   479
          Drule.instantiate_normalize ([], [(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
   480
      in
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   481
        Conv.rewr_conv thm4 ctrm
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   482
      end
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   483
  | _ => Conv.all_conv ctrm)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   484
36936
c52d1c130898 incorporated further conversions and conversionals, after some minor tuning;
wenzelm
parents: 36850
diff changeset
   485
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
   486
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
   487
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
(* Cleaning consists of:
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   490
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   491
  1. unfolding of ---> in front of everything, except
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   492
     bound variables (this prevents lambda_prs from
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   493
     becoming stuck)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   494
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   495
  2. simplification with lambda_prs
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   496
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   497
  3. simplification with:
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   498
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   499
      - Quotient_abs_rep Quotient_rel_rep
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   500
        babs_prs all_prs ex_prs ex1_prs
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   501
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   502
      - id_simps and preservation lemmas and
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   503
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   504
      - symmetric versions of the definitions
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   505
        (that is definitions of quotient constants
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   506
         are folded)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   507
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   508
  4. test for refl
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   509
*)
57960
ee1ba4848896 updated to named_theorems;
wenzelm
parents: 55945
diff changeset
   510
fun clean_tac ctxt =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   511
  let
57960
ee1ba4848896 updated to named_theorems;
wenzelm
parents: 55945
diff changeset
   512
    val thy =  Proof_Context.theory_of ctxt
45350
257d0b179f0d more use of global operations (see 98ec8b51af9c)
Christian Urban <urbanc@in.tum.de>
parents: 45279
diff changeset
   513
    val defs = map (Thm.symmetric o #def) (Quotient_Info.dest_quotconsts_global thy)
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   514
    val prs = rev (Named_Theorems.get ctxt \<^named_theorems>\<open>quot_preserve\<close>)
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67632
diff changeset
   515
    val ids = rev (Named_Theorems.get ctxt \<^named_theorems>\<open>id_simps\<close>)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   516
    val thms =
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents: 46468
diff changeset
   517
      @{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
   518
57960
ee1ba4848896 updated to named_theorems;
wenzelm
parents: 55945
diff changeset
   519
    val simpset = (mk_minimal_simpset ctxt) addsimps thms addSolver quotient_solver
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   520
  in
41451
892e67be8304 do not open ML structures;
wenzelm
parents: 41444
diff changeset
   521
    EVERY' [
57960
ee1ba4848896 updated to named_theorems;
wenzelm
parents: 55945
diff changeset
   522
      map_fun_tac ctxt,
ee1ba4848896 updated to named_theorems;
wenzelm
parents: 55945
diff changeset
   523
      lambda_prs_tac ctxt,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47308
diff changeset
   524
      simp_tac simpset,
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   525
      TRY o resolve_tac ctxt [refl]]
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   526
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   527
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   528
38718
c7cbbb18eabe tuned code
Christian Urban <urbanc@in.tum.de>
parents: 38717
diff changeset
   529
(* 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
   530
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   531
fun inst_spec ctrm =
60801
7664e0916eec tuned signature;
wenzelm
parents: 60752
diff changeset
   532
  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
   533
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   534
fun inst_spec_tac ctxt ctrms =
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   535
  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
   536
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   537
fun all_list xs trm =
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   538
  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
   539
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   540
fun apply_under_Trueprop f =
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   541
  HOLogic.dest_Trueprop #> f #> HOLogic.mk_Trueprop
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   542
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   543
fun gen_frees_tac ctxt =
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   544
  SUBGOAL (fn (concl, i) =>
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   545
    let
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   546
      val vrs = Term.add_frees concl []
59638
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   547
      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
   548
      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
   549
      val goal = Logic.mk_implies (concl', concl)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   550
      val rule = Goal.prove ctxt [] [] goal
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   551
        (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
   552
    in
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   553
      resolve_tac ctxt [rule] i
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   554
    end)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   555
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
(** The General Shape of the Lifting Procedure **)
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   558
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   559
(* - A is the original raw theorem
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   560
   - B is the regularized theorem
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   561
   - 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
   562
   - D is the lifted theorem
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   563
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   564
   - 1st prem is the regularization step
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   565
   - 2nd prem is the rep/abs injection step
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   566
   - 3rd prem is the cleaning part
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   567
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   568
   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
   569
*)
45782
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   570
val procedure_thm =
67091
1393c2340eec more symbols;
wenzelm
parents: 62913
diff changeset
   571
  @{lemma  "\<lbrakk>A;
1393c2340eec more symbols;
wenzelm
parents: 62913
diff changeset
   572
              A \<longrightarrow> B;
1393c2340eec more symbols;
wenzelm
parents: 62913
diff changeset
   573
              Quot_True D \<Longrightarrow> B = C;
1393c2340eec more symbols;
wenzelm
parents: 62913
diff changeset
   574
              C = D\<rbrakk> \<Longrightarrow> D"
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   575
      by (simp add: Quot_True_def)}
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   576
45782
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   577
(* 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: 45625
diff changeset
   578
   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: 45625
diff changeset
   579
*)
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   580
val partiality_procedure_thm =
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   581
  @{lemma  "[|B; 
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   582
              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: 45625
diff changeset
   583
              C = D|] ==> D"
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   584
      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: 45625
diff changeset
   585
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   586
fun lift_match_error ctxt msg rtrm qtrm =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   587
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   588
    val rtrm_str = Syntax.string_of_term ctxt rtrm
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   589
    val qtrm_str = Syntax.string_of_term ctxt qtrm
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   590
    val msg = cat_lines [enclose "[" "]" msg, "The quotient theorem", qtrm_str,
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   591
      "", "does not match with original theorem", rtrm_str]
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   592
  in
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   593
    error msg
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   594
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   595
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   596
fun procedure_inst ctxt rtrm qtrm =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   597
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   598
    val rtrm' = HOLogic.dest_Trueprop rtrm
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   599
    val qtrm' = HOLogic.dest_Trueprop qtrm
41451
892e67be8304 do not open ML structures;
wenzelm
parents: 41444
diff changeset
   600
    val reg_goal = Quotient_Term.regularize_trm_chk ctxt (rtrm', qtrm')
892e67be8304 do not open ML structures;
wenzelm
parents: 41444
diff changeset
   601
      handle Quotient_Term.LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm
892e67be8304 do not open ML structures;
wenzelm
parents: 41444
diff changeset
   602
    val inj_goal = Quotient_Term.inj_repabs_trm_chk ctxt (reg_goal, qtrm')
892e67be8304 do not open ML structures;
wenzelm
parents: 41444
diff changeset
   603
      handle Quotient_Term.LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   604
  in
60801
7664e0916eec tuned signature;
wenzelm
parents: 60752
diff changeset
   605
    Thm.instantiate' []
59638
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   606
      [SOME (Thm.cterm_of ctxt rtrm'),
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   607
       SOME (Thm.cterm_of ctxt reg_goal),
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   608
       NONE,
59638
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   609
       SOME (Thm.cterm_of ctxt inj_goal)] procedure_thm
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   610
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   611
37593
2505feaf2d70 separated the lifting and descending procedures in the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 37592
diff changeset
   612
38860
749d09f52fde quotient package: added a list of pre-simplification rules for Ball, Bex and mem
Christian Urban <urbanc@in.tum.de>
parents: 38859
diff changeset
   613
(* 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: 38860
diff changeset
   614
   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: 38860
diff changeset
   615
   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: 38859
diff changeset
   616
38862
2795499a20bd Quotient Package: dont unfold mem_def, use rsp and prs instead
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 38860
diff changeset
   617
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: 38859
diff changeset
   618
749d09f52fde quotient package: added a list of pre-simplification rules for Ball, Bex and mem
Christian Urban <urbanc@in.tum.de>
parents: 38859
diff changeset
   619
37593
2505feaf2d70 separated the lifting and descending procedures in the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 37592
diff changeset
   620
(** descending as tactic **)
2505feaf2d70 separated the lifting and descending procedures in the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 37592
diff changeset
   621
38859
053c69cb4a0e quotient package: lemmas to be lifted and descended can be pre-simplified
Christian Urban <urbanc@in.tum.de>
parents: 38719
diff changeset
   622
fun descend_procedure_tac ctxt simps =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   623
  let
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47308
diff changeset
   624
    val simpset = (mk_minimal_simpset ctxt) addsimps (simps @ default_unfolds)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   625
  in
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47308
diff changeset
   626
    full_simp_tac simpset
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52230
diff changeset
   627
    THEN' Object_Logic.full_atomize_tac ctxt
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   628
    THEN' gen_frees_tac ctxt
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   629
    THEN' SUBGOAL (fn (goal, i) =>
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   630
      let
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 44285
diff changeset
   631
        val qtys = map #qtyp (Quotient_Info.dest_quotients ctxt)
41451
892e67be8304 do not open ML structures;
wenzelm
parents: 41444
diff changeset
   632
        val rtrm = Quotient_Term.derive_rtrm ctxt qtys goal
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   633
        val rule = procedure_inst ctxt rtrm  goal
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   634
      in
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   635
        resolve_tac ctxt [rule] i
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   636
      end)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   637
  end
37593
2505feaf2d70 separated the lifting and descending procedures in the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 37592
diff changeset
   638
38859
053c69cb4a0e quotient package: lemmas to be lifted and descended can be pre-simplified
Christian Urban <urbanc@in.tum.de>
parents: 38719
diff changeset
   639
fun descend_tac ctxt simps =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   640
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   641
    val mk_tac_raw =
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   642
      descend_procedure_tac ctxt simps
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   643
      THEN' RANGE
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52230
diff changeset
   644
        [Object_Logic.rulify_tac ctxt THEN' (K all_tac),
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   645
         regularize_tac ctxt,
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   646
         all_injection_tac ctxt,
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   647
         clean_tac ctxt]
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   648
  in
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   649
    Goal.conjunction_tac THEN_ALL_NEW mk_tac_raw
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   650
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   651
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   652
45782
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   653
(** 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: 45625
diff changeset
   654
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   655
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: 45625
diff changeset
   656
  let
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   657
    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: 45625
diff changeset
   658
    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: 45625
diff changeset
   659
    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: 45625
diff changeset
   660
      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: 45625
diff changeset
   661
    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: 45625
diff changeset
   662
      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: 45625
diff changeset
   663
  in
60801
7664e0916eec tuned signature;
wenzelm
parents: 60752
diff changeset
   664
    Thm.instantiate' []
59638
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   665
      [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: 45625
diff changeset
   666
       NONE,
59638
cb84e420fc8e clarified context;
wenzelm
parents: 59621
diff changeset
   667
       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: 45625
diff changeset
   668
  end
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   669
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   670
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   671
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: 45625
diff changeset
   672
  let
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47308
diff changeset
   673
    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: 45625
diff changeset
   674
  in
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47308
diff changeset
   675
    full_simp_tac simpset
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52230
diff changeset
   676
    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: 45625
diff changeset
   677
    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: 45625
diff changeset
   678
    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: 45625
diff changeset
   679
      let
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   680
        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: 45625
diff changeset
   681
        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: 45625
diff changeset
   682
        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: 45625
diff changeset
   683
      in
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   684
        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: 45625
diff changeset
   685
      end)
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   686
  end
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   687
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   688
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: 45625
diff changeset
   689
  let
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   690
    val mk_tac_raw =
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   691
      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: 45625
diff changeset
   692
      THEN' RANGE
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52230
diff changeset
   693
        [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: 45625
diff changeset
   694
         all_injection_tac ctxt,
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   695
         clean_tac ctxt]
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   696
  in
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   697
    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: 45625
diff changeset
   698
  end
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   699
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   700
f82020ca3248 added a specific tactic and method that deal with partial equivalence relations
Christian Urban <urbanc@in.tum.de>
parents: 45625
diff changeset
   701
38625
b97bd93fb9e2 allow for pre-simplification of lifted theorems
Christian Urban <urbanc@in.tum.de>
parents: 38624
diff changeset
   702
(** lifting as a tactic **)
37593
2505feaf2d70 separated the lifting and descending procedures in the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 37592
diff changeset
   703
38718
c7cbbb18eabe tuned code
Christian Urban <urbanc@in.tum.de>
parents: 38717
diff changeset
   704
37593
2505feaf2d70 separated the lifting and descending procedures in the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 37592
diff changeset
   705
(* 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: 38719
diff changeset
   706
fun lift_procedure_tac ctxt simps rthm =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   707
  let
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47308
diff changeset
   708
    val simpset = (mk_minimal_simpset ctxt) addsimps (simps @ default_unfolds)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   709
  in
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47308
diff changeset
   710
    full_simp_tac simpset
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52230
diff changeset
   711
    THEN' Object_Logic.full_atomize_tac ctxt
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   712
    THEN' gen_frees_tac ctxt
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   713
    THEN' SUBGOAL (fn (goal, i) =>
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   714
      let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   715
        (* full_atomize_tac contracts eta redexes,
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   716
           so we do it also in the original theorem *)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   717
        val rthm' =
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47308
diff changeset
   718
          rthm |> full_simplify simpset
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   719
               |> Drule.eta_contraction_rule
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   720
               |> Thm.forall_intr_frees
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52230
diff changeset
   721
               |> atomize_thm ctxt
38717
a365f1fc5081 quotient package: deal correctly with frees in lifted theorems
Christian Urban <urbanc@in.tum.de>
parents: 38694
diff changeset
   722
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   723
        val rule = procedure_inst ctxt (Thm.prop_of rthm') goal
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   724
      in
60752
b48830b670a1 prefer tactics with explicit context;
wenzelm
parents: 60642
diff changeset
   725
        (resolve_tac ctxt [rule] THEN' resolve_tac ctxt [rthm']) i
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   726
      end)
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   727
  end
38625
b97bd93fb9e2 allow for pre-simplification of lifted theorems
Christian Urban <urbanc@in.tum.de>
parents: 38624
diff changeset
   728
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   729
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: 38719
diff changeset
   730
  lift_procedure_tac ctxt simps rthm
38625
b97bd93fb9e2 allow for pre-simplification of lifted theorems
Christian Urban <urbanc@in.tum.de>
parents: 38624
diff changeset
   731
  THEN' RANGE
b97bd93fb9e2 allow for pre-simplification of lifted theorems
Christian Urban <urbanc@in.tum.de>
parents: 38624
diff changeset
   732
    [ regularize_tac ctxt,
b97bd93fb9e2 allow for pre-simplification of lifted theorems
Christian Urban <urbanc@in.tum.de>
parents: 38624
diff changeset
   733
      all_injection_tac ctxt,
b97bd93fb9e2 allow for pre-simplification of lifted theorems
Christian Urban <urbanc@in.tum.de>
parents: 38624
diff changeset
   734
      clean_tac ctxt ]
b97bd93fb9e2 allow for pre-simplification of lifted theorems
Christian Urban <urbanc@in.tum.de>
parents: 38624
diff changeset
   735
38859
053c69cb4a0e quotient package: lemmas to be lifted and descended can be pre-simplified
Christian Urban <urbanc@in.tum.de>
parents: 38719
diff changeset
   736
fun lift_tac ctxt simps rthms =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   737
  Goal.conjunction_tac
38859
053c69cb4a0e quotient package: lemmas to be lifted and descended can be pre-simplified
Christian Urban <urbanc@in.tum.de>
parents: 38719
diff changeset
   738
  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: 38624
diff changeset
   739
b97bd93fb9e2 allow for pre-simplification of lifted theorems
Christian Urban <urbanc@in.tum.de>
parents: 38624
diff changeset
   740
b97bd93fb9e2 allow for pre-simplification of lifted theorems
Christian Urban <urbanc@in.tum.de>
parents: 38624
diff changeset
   741
(* automated lifting with pre-simplification of the theorems;
b97bd93fb9e2 allow for pre-simplification of lifted theorems
Christian Urban <urbanc@in.tum.de>
parents: 38624
diff changeset
   742
   for internal usage *)
b97bd93fb9e2 allow for pre-simplification of lifted theorems
Christian Urban <urbanc@in.tum.de>
parents: 38624
diff changeset
   743
fun lifted ctxt qtys simps rthm =
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   744
  let
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   745
    val ((_, [rthm']), ctxt') = Variable.import true [rthm] ctxt
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   746
    val goal = Quotient_Term.derive_qtrm ctxt' qtys (Thm.prop_of rthm')
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   747
  in
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   748
    Goal.prove ctxt' [] [] goal
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   749
      (K (HEADGOAL (lift_single_tac ctxt' simps rthm')))
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 41451
diff changeset
   750
    |> singleton (Proof_Context.export ctxt' ctxt)
41444
7f40120cd814 more precise parentheses and indentation;
wenzelm
parents: 41443
diff changeset
   751
  end
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   752
37593
2505feaf2d70 separated the lifting and descending procedures in the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 37592
diff changeset
   753
38625
b97bd93fb9e2 allow for pre-simplification of lifted theorems
Christian Urban <urbanc@in.tum.de>
parents: 38624
diff changeset
   754
(* lifting as an attribute *)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   755
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: 61144
diff changeset
   756
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: 37592
diff changeset
   757
  let
2505feaf2d70 separated the lifting and descending procedures in the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 37592
diff changeset
   758
    val ctxt = Context.proof_of context
45279
89a17197cb98 simplified/standardized signatures;
wenzelm
parents: 44285
diff changeset
   759
    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: 37592
diff changeset
   760
  in
38625
b97bd93fb9e2 allow for pre-simplification of lifted theorems
Christian Urban <urbanc@in.tum.de>
parents: 38624
diff changeset
   761
    lifted ctxt qtys []
37593
2505feaf2d70 separated the lifting and descending procedures in the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 37592
diff changeset
   762
  end)
35222
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   763
4f1fba00f66d Initial version of HOL quotient package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff changeset
   764
end; (* structure *)