Attic/Quot/quotient_tacs.ML
author Christian Urban <urbanc@in.tum.de>
Fri, 14 May 2010 17:40:43 +0100
changeset 2136 2fc55508a6d0
parent 1450 1ae5afcddcd4
permissions -rw-r--r--
polished example
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1438
61671de8a545 synchronised with main hg-repository; used add_typedef_global in nominal_atoms
Christian Urban <urbanc@in.tum.de>
parents: 1354
diff changeset
     1
(*  Title:      HOL/Tools/Quotient/quotient_tacs.thy
952
9c3b3eaecaff use of equiv_relation_chk in quotient_term
Christian Urban <urbanc@in.tum.de>
parents: 951
diff changeset
     2
    Author:     Cezary Kaliszyk and Christian Urban
9c3b3eaecaff use of equiv_relation_chk in quotient_term
Christian Urban <urbanc@in.tum.de>
parents: 951
diff changeset
     3
1438
61671de8a545 synchronised with main hg-repository; used add_typedef_global in nominal_atoms
Christian Urban <urbanc@in.tum.de>
parents: 1354
diff changeset
     4
Tactics for solving goal arising from lifting theorems to quotient
61671de8a545 synchronised with main hg-repository; used add_typedef_global in nominal_atoms
Christian Urban <urbanc@in.tum.de>
parents: 1354
diff changeset
     5
types.
952
9c3b3eaecaff use of equiv_relation_chk in quotient_term
Christian Urban <urbanc@in.tum.de>
parents: 951
diff changeset
     6
*)
9c3b3eaecaff use of equiv_relation_chk in quotient_term
Christian Urban <urbanc@in.tum.de>
parents: 951
diff changeset
     7
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
signature QUOTIENT_TACS =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
sig
853
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
    10
  val regularize_tac: Proof.context -> int -> tactic
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
    11
  val injection_tac: Proof.context -> int -> tactic
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
    12
  val all_injection_tac: Proof.context -> int -> tactic
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
    13
  val clean_tac: Proof.context -> int -> tactic
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
    14
  val procedure_tac: Proof.context -> thm -> int -> tactic
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 871
diff changeset
    15
  val lift_tac: Proof.context -> thm list -> int -> tactic
853
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
    16
  val quotient_tac: Proof.context -> int -> tactic
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
    17
  val quot_true_tac: Proof.context -> (term -> term) -> int -> tactic
1077
44461d5615eb some tuning
Christian Urban <urbanc@in.tum.de>
parents: 1075
diff changeset
    18
  val lifted_attrib: attribute
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
end;
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
structure Quotient_Tacs: QUOTIENT_TACS =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
struct
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
762
baac4639ecef avoided global "open"s - replaced by local "open"s
Christian Urban <urbanc@in.tum.de>
parents: 761
diff changeset
    24
open Quotient_Info;
baac4639ecef avoided global "open"s - replaced by local "open"s
Christian Urban <urbanc@in.tum.de>
parents: 761
diff changeset
    25
open Quotient_Term;
baac4639ecef avoided global "open"s - replaced by local "open"s
Christian Urban <urbanc@in.tum.de>
parents: 761
diff changeset
    26
848
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
    27
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
    28
(** various helper fuctions **)
802
Christian Urban <urbanc@in.tum.de>
parents: 789
diff changeset
    29
769
d89851ebac9b cleaned a bit calculate_inst a bit; eta-contraction seems to be not necessary? (all examples go through)
Christian Urban <urbanc@in.tum.de>
parents: 768
diff changeset
    30
(* Since HOL_basic_ss is too "big" for us, we *)
d89851ebac9b cleaned a bit calculate_inst a bit; eta-contraction seems to be not necessary? (all examples go through)
Christian Urban <urbanc@in.tum.de>
parents: 768
diff changeset
    31
(* need to set up our own minimal simpset.    *)
1113
9f6c606d5b59 more minor space and bracket modifications.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1112
diff changeset
    32
fun mk_minimal_ss ctxt =
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
  Simplifier.context ctxt empty_ss
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
    setsubgoaler asm_simp_tac
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
    setmksimps (mksimps [])
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
785
Christian Urban <urbanc@in.tum.de>
parents: 781
diff changeset
    37
(* composition of two theorems, used in maps *)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
fun OF1 thm1 thm2 = thm2 RS thm1
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
802
Christian Urban <urbanc@in.tum.de>
parents: 789
diff changeset
    40
(* prints a warning, if the subgoal is not solved *)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
fun WARN (tac, msg) i st =
1112
c7069b09730b More changes according to the standards.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1102
diff changeset
    42
 case Seq.pull (SOLVED' tac i st) of
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
     NONE    => (warning msg; Seq.single st)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
   | seqcell => Seq.make (fn () => seqcell)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
802
Christian Urban <urbanc@in.tum.de>
parents: 789
diff changeset
    46
fun RANGE_WARN tacs = RANGE (map WARN tacs)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
fun atomize_thm thm =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
let
804
Christian Urban <urbanc@in.tum.de>
parents: 802
diff changeset
    50
  val thm' = Thm.freezeT (forall_intr_vars thm) (* FIXME/TODO: is this proper Isar-technology? *)
1354
367f67311e6f updated to renamings in Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 1260
diff changeset
    51
  val thm'' = Object_Logic.atomize (cprop_of thm')
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
  @{thm equal_elim_rule1} OF [thm'', thm']
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
848
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
    57
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
    58
(*** Regularize Tactic ***)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
848
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
    60
(** solvers for equivp and quotient assumptions **)
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
    61
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
    62
fun equiv_tac ctxt =
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
  REPEAT_ALL_NEW (resolve_tac (equiv_rules_get ctxt))
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
fun equiv_solver_tac ss = equiv_tac (Simplifier.the_context ss)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
val equiv_solver = Simplifier.mk_solver' "Equivalence goal solver" equiv_solver_tac
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
870
2a19e0a37131 Remove SOLVED from quotient_tac. Move atomize_eqv to 'Unused'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 866
diff changeset
    68
fun quotient_tac ctxt =
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
    69
  (REPEAT_ALL_NEW (FIRST'
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
    70
    [rtac @{thm identity_quotient},
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
    71
     resolve_tac (quotient_rules_get ctxt)]))
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
    72
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
    73
fun quotient_solver_tac ss = quotient_tac (Simplifier.the_context ss)
853
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
    74
val quotient_solver =
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
    75
  Simplifier.mk_solver' "Quotient goal solver" quotient_solver_tac
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
    76
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
    77
fun solve_quotient_assm ctxt thm =
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
    78
  case Seq.pull (quotient_tac ctxt 1 thm) of
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
    79
    SOME (t, _) => t
802
Christian Urban <urbanc@in.tum.de>
parents: 789
diff changeset
    80
  | _ => error "Solve_quotient_assm failed. Possibly a quotient theorem is missing."
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
    81
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
    82
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
fun prep_trm thy (x, (T, t)) =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
  (cterm_of thy (Var (x, T)), cterm_of thy t)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
fun prep_ty thy (x, (S, ty)) =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
  (ctyp_of thy (TVar (x, S)), ctyp_of thy ty)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
772
a95f6bb081cf on the hunt for what condition raises which exception in the CLEVER CODE of calculate_inst
Christian Urban <urbanc@in.tum.de>
parents: 771
diff changeset
    89
fun get_match_inst thy pat trm =
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
  val univ = Unify.matchers thy [(pat, trm)]
1438
61671de8a545 synchronised with main hg-repository; used add_typedef_global in nominal_atoms
Christian Urban <urbanc@in.tum.de>
parents: 1354
diff changeset
    92
  val SOME (env, _) = Seq.pull univ           (* raises Bind, if no unifier *)  (* FIXME fragile *)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
  val tenv = Vartab.dest (Envir.term_env env)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
  val tyenv = Vartab.dest (Envir.type_env env)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
  (map (prep_ty thy) tyenv, map (prep_trm thy) tenv)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
837
116c7a30e0a2 Changing exceptions to 'try', part 1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 828
diff changeset
    99
(* Calculates the instantiations for the lemmas:
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   100
837
116c7a30e0a2 Changing exceptions to 'try', part 1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 828
diff changeset
   101
      ball_reg_eqv_range and bex_reg_eqv_range
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   102
837
116c7a30e0a2 Changing exceptions to 'try', part 1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 828
diff changeset
   103
   Since the left-hand-side contains a non-pattern '?P (f ?x)'
116c7a30e0a2 Changing exceptions to 'try', part 1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 828
diff changeset
   104
   we rely on unification/instantiation to check whether the
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   105
   theorem applies and return NONE if it doesn't.
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   106
*)
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   107
fun calculate_inst ctxt ball_bex_thm redex R1 R2 =
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
let
951
62f0344b219c some slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 946
diff changeset
   109
  val thy = ProofContext.theory_of ctxt
771
b2231990b059 simplified calculate_instance; worked around some clever code; clever code is unfortunately still there...needs to be removed
Christian Urban <urbanc@in.tum.de>
parents: 770
diff changeset
   110
  fun get_lhs thm = fst (Logic.dest_equals (Thm.concl_of thm))
951
62f0344b219c some slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 946
diff changeset
   111
  val ty_inst = map (SOME o ctyp_of thy) [domain_type (fastype_of R2)]
62f0344b219c some slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 946
diff changeset
   112
  val trm_inst = map (SOME o cterm_of thy) [R2, R1]
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
in
951
62f0344b219c some slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 946
diff changeset
   114
  case try (Drule.instantiate' ty_inst trm_inst) ball_bex_thm of
837
116c7a30e0a2 Changing exceptions to 'try', part 1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 828
diff changeset
   115
    NONE => NONE
116c7a30e0a2 Changing exceptions to 'try', part 1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 828
diff changeset
   116
  | SOME thm' =>
116c7a30e0a2 Changing exceptions to 'try', part 1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 828
diff changeset
   117
      (case try (get_match_inst thy (get_lhs thm')) redex of
116c7a30e0a2 Changing exceptions to 'try', part 1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 828
diff changeset
   118
        NONE => NONE
116c7a30e0a2 Changing exceptions to 'try', part 1.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 828
diff changeset
   119
      | SOME inst2 => try (Drule.instantiate inst2) thm')
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
fun ball_bex_range_simproc ss redex =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
  val ctxt = Simplifier.the_context ss
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   125
in
772
a95f6bb081cf on the hunt for what condition raises which exception in the CLEVER CODE of calculate_inst
Christian Urban <urbanc@in.tum.de>
parents: 771
diff changeset
   126
  case redex of
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   127
    (Const (@{const_name "Ball"}, _) $ (Const (@{const_name "Respects"}, _) $
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
      (Const (@{const_name "fun_rel"}, _) $ R1 $ R2)) $ _) =>
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   129
        calculate_inst ctxt @{thm ball_reg_eqv_range[THEN eq_reflection]} redex R1 R2
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   131
  | (Const (@{const_name "Bex"}, _) $ (Const (@{const_name "Respects"}, _) $
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   132
      (Const (@{const_name "fun_rel"}, _) $ R1 $ R2)) $ _) =>
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   133
        calculate_inst ctxt @{thm bex_reg_eqv_range[THEN eq_reflection]} redex R1 R2
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   134
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
  | _ => NONE
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   138
(* Regularize works as follows:
853
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   139
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   140
  0. preliminary simplification step according to
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   141
     ball_reg_eqv bex_reg_eqv babs_reg_eqv ball_reg_eqv_range bex_reg_eqv_range
853
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   142
1074
7a42cc191111 Fixes for Bex1 removal.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 980
diff changeset
   143
  1. eliminating simple Ball/Bex instances (ball_reg_right bex_reg_left)
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   144
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   145
  2. monos
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   147
  3. commutation rules for ball and bex (ball_all_comm bex_ex_comm)
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   148
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   149
  4. then rel-equalities, which need to be instantiated with 'eq_imp_rel'
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   150
     to avoid loops
853
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   151
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   152
  5. then simplification like 0
853
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   153
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   154
  finally jump back to 1
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   155
*)
1137
36d596cb63a2 moved "strange" lemma to quotient_tacs; marked a number of lemmas as unused; tuned
Christian Urban <urbanc@in.tum.de>
parents: 1128
diff changeset
   156
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
fun regularize_tac ctxt =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
  val thy = ProofContext.theory_of ctxt
772
a95f6bb081cf on the hunt for what condition raises which exception in the CLEVER CODE of calculate_inst
Christian Urban <urbanc@in.tum.de>
parents: 771
diff changeset
   160
  val ball_pat = @{term "Ball (Respects (R1 ===> R2)) P"}
a95f6bb081cf on the hunt for what condition raises which exception in the CLEVER CODE of calculate_inst
Christian Urban <urbanc@in.tum.de>
parents: 771
diff changeset
   161
  val bex_pat  = @{term "Bex (Respects (R1 ===> R2)) P"}
a95f6bb081cf on the hunt for what condition raises which exception in the CLEVER CODE of calculate_inst
Christian Urban <urbanc@in.tum.de>
parents: 771
diff changeset
   162
  val simproc = Simplifier.simproc_i thy "" [ball_pat, bex_pat] (K (ball_bex_range_simproc))
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   163
  val simpset = (mk_minimal_ss ctxt)
845
9531fafc0daa removed 3 hacks.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 844
diff changeset
   164
                       addsimps @{thms ball_reg_eqv bex_reg_eqv babs_reg_eqv babs_simp}
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   165
                       addsimprocs [simproc]
772
a95f6bb081cf on the hunt for what condition raises which exception in the CLEVER CODE of calculate_inst
Christian Urban <urbanc@in.tum.de>
parents: 771
diff changeset
   166
                       addSolver equiv_solver addSolver quotient_solver
1137
36d596cb63a2 moved "strange" lemma to quotient_tacs; marked a number of lemmas as unused; tuned
Christian Urban <urbanc@in.tum.de>
parents: 1128
diff changeset
   167
  val eq_imp_rel = @{lemma "equivp R ==> a = b --> R a b" by (simp add: equivp_reflp)}
36d596cb63a2 moved "strange" lemma to quotient_tacs; marked a number of lemmas as unused; tuned
Christian Urban <urbanc@in.tum.de>
parents: 1128
diff changeset
   168
  val eq_eqvs = map (OF1 eq_imp_rel) (equiv_rules_get ctxt)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
  simp_tac simpset THEN'
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   171
  REPEAT_ALL_NEW (CHANGED o FIRST'
946
46cc6708c3b3 Bex1_Bexeq_regular.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 940
diff changeset
   172
    [resolve_tac @{thms ball_reg_right bex_reg_left bex1_bexeq_reg},
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   173
     resolve_tac (Inductive.get_monos ctxt),
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   174
     resolve_tac @{thms ball_all_comm bex_ex_comm},
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   175
     resolve_tac eq_eqvs,
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   176
     simp_tac simpset])
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
848
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   180
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   181
(*** Injection Tactic ***)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
1112
c7069b09730b More changes according to the standards.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1102
diff changeset
   183
(* Looks for Quot_True assumptions, and in case its parameter
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   184
   is an application, it returns the function and the argument.
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   185
*)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
fun find_qt_asm asms =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
  fun find_fun trm =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
    case trm of
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   190
      (Const(@{const_name Trueprop}, _) $ (Const (@{const_name Quot_True}, _) $ _)) => true
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
    | _ => false
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
 case find_first find_fun asms of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
   SOME (_ $ (_ $ (f $ a))) => SOME (f, a)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
 | _ => NONE
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
fun quot_true_simple_conv ctxt fnctn ctrm =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
  case (term_of ctrm) of
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   200
    (Const (@{const_name Quot_True}, _) $ x) =>
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
    let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
      val fx = fnctn x;
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
      val thy = ProofContext.theory_of ctxt;
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
      val cx = cterm_of thy x;
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
      val cfx = cterm_of thy fx;
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
      val cxt = ctyp_of thy (fastype_of x);
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
      val cfxt = ctyp_of thy (fastype_of fx);
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   208
      val thm = Drule.instantiate' [SOME cxt, SOME cfxt] [SOME cx, SOME cfx] @{thm QT_imp}
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
    in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
      Conv.rewr_conv thm ctrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
    end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
fun quot_true_conv ctxt fnctn ctrm =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
  case (term_of ctrm) of
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   215
    (Const (@{const_name Quot_True}, _) $ _) =>
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
      quot_true_simple_conv ctxt fnctn ctrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
  | _ $ _ => Conv.comb_conv (quot_true_conv ctxt fnctn) ctrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
  | Abs _ => Conv.abs_conv (fn (_, ctxt) => quot_true_conv ctxt fnctn) ctxt ctrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
  | _ => Conv.all_conv ctrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   221
fun quot_true_tac ctxt fnctn =
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
   CONVERSION
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
    ((Conv.params_conv ~1 (fn ctxt =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
       (Conv.prems_conv ~1 (quot_true_conv ctxt fnctn)))) ctxt)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   226
fun dest_comb (f $ a) = (f, a)
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   227
fun dest_bcomb ((_ $ l) $ r) = (l, r)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
fun unlam t =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
  case t of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
    (Abs a) => snd (Term.dest_abs a)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
  | _ => unlam (Abs("", domain_type (fastype_of t), (incr_boundvars 1 t) $ (Bound 0)))
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
fun dest_fun_type (Type("fun", [T, S])) = (T, S)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
  | dest_fun_type _ = error "dest_fun_type"
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
val bare_concl = HOLogic.dest_Trueprop o Logic.strip_assums_concl
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
848
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   239
(* We apply apply_rsp only in case if the type needs lifting.
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   240
   This is the case if the type of the data in the Quot_True
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   241
   assumption is different from the corresponding type in the goal.
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   242
*)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
val apply_rsp_tac =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
  Subgoal.FOCUS (fn {concl, asms, context,...} =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
  let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
    val bare_concl = HOLogic.dest_Trueprop (term_of concl)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
    val qt_asm = find_qt_asm (map term_of asms)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
  in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
    case (bare_concl, qt_asm) of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
      (R2 $ (f $ x) $ (g $ y), SOME (qt_fun, qt_arg)) =>
1102
26f15c2dc131 removing unnecessary brackets
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1084
diff changeset
   251
         if fastype_of qt_fun = fastype_of f
848
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   252
         then no_tac
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   253
         else
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
           let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
             val ty_x = fastype_of x
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
             val ty_b = fastype_of qt_arg
848
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   257
             val ty_f = range_type (fastype_of f)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
             val thy = ProofContext.theory_of context
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
             val ty_inst = map (SOME o (ctyp_of thy)) [ty_x, ty_b, ty_f]
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
             val t_inst = map (SOME o (cterm_of thy)) [R2, f, g, x, y];
848
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   261
             val inst_thm = Drule.instantiate' ty_inst
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   262
               ([NONE, NONE, NONE] @ t_inst) @{thm apply_rsp}
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
           in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
             (rtac inst_thm THEN' quotient_tac context) 1
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
           end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
    | _ => no_tac
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
  end)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
842
0332d0df2fc9 Removed exception handling from equals_rsp_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 839
diff changeset
   269
(* Instantiates and applies 'equals_rsp'. Since the theorem is
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   270
   complex we rely on instantiation to tell us if it applies
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   271
*)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
fun equals_rsp_tac R ctxt =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
  val thy = ProofContext.theory_of ctxt
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
in
842
0332d0df2fc9 Removed exception handling from equals_rsp_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 839
diff changeset
   276
  case try (cterm_of thy) R of (* There can be loose bounds in R *)
0332d0df2fc9 Removed exception handling from equals_rsp_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 839
diff changeset
   277
    SOME ctm =>
0332d0df2fc9 Removed exception handling from equals_rsp_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 839
diff changeset
   278
      let
0332d0df2fc9 Removed exception handling from equals_rsp_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 839
diff changeset
   279
        val ty = domain_type (fastype_of R)
0332d0df2fc9 Removed exception handling from equals_rsp_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 839
diff changeset
   280
      in
0332d0df2fc9 Removed exception handling from equals_rsp_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 839
diff changeset
   281
        case try (Drule.instantiate' [SOME (ctyp_of thy ty)]
0332d0df2fc9 Removed exception handling from equals_rsp_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 839
diff changeset
   282
          [SOME (cterm_of thy R)]) @{thm equals_rsp} of
0332d0df2fc9 Removed exception handling from equals_rsp_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 839
diff changeset
   283
          SOME thm => rtac thm THEN' quotient_tac ctxt
0332d0df2fc9 Removed exception handling from equals_rsp_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 839
diff changeset
   284
        | NONE => K no_tac
0332d0df2fc9 Removed exception handling from equals_rsp_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 839
diff changeset
   285
      end
0332d0df2fc9 Removed exception handling from equals_rsp_tac.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 839
diff changeset
   286
  | _ => K no_tac
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   287
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
847
b89707cd030f No more exception handling in rep_abs_rsp_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 846
diff changeset
   289
fun rep_abs_rsp_tac ctxt =
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
  SUBGOAL (fn (goal, i) =>
847
b89707cd030f No more exception handling in rep_abs_rsp_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 846
diff changeset
   291
    case (try bare_concl goal) of
b89707cd030f No more exception handling in rep_abs_rsp_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 846
diff changeset
   292
      SOME (rel $ _ $ (rep $ (abs $ _))) =>
b89707cd030f No more exception handling in rep_abs_rsp_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 846
diff changeset
   293
        let
b89707cd030f No more exception handling in rep_abs_rsp_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 846
diff changeset
   294
          val thy = ProofContext.theory_of ctxt;
b89707cd030f No more exception handling in rep_abs_rsp_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 846
diff changeset
   295
          val (ty_a, ty_b) = dest_fun_type (fastype_of abs);
b89707cd030f No more exception handling in rep_abs_rsp_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 846
diff changeset
   296
          val ty_inst = map (SOME o (ctyp_of thy)) [ty_a, ty_b];
b89707cd030f No more exception handling in rep_abs_rsp_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 846
diff changeset
   297
        in
b89707cd030f No more exception handling in rep_abs_rsp_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 846
diff changeset
   298
          case try (map (SOME o (cterm_of thy))) [rel, abs, rep] of
b89707cd030f No more exception handling in rep_abs_rsp_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 846
diff changeset
   299
            SOME t_inst =>
b89707cd030f No more exception handling in rep_abs_rsp_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 846
diff changeset
   300
              (case try (Drule.instantiate' ty_inst t_inst) @{thm rep_abs_rsp} of
b89707cd030f No more exception handling in rep_abs_rsp_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 846
diff changeset
   301
                SOME inst_thm => (rtac inst_thm THEN' quotient_tac ctxt) i
b89707cd030f No more exception handling in rep_abs_rsp_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 846
diff changeset
   302
              | NONE => no_tac)
b89707cd030f No more exception handling in rep_abs_rsp_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 846
diff changeset
   303
          | NONE => no_tac
b89707cd030f No more exception handling in rep_abs_rsp_tac
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 846
diff changeset
   304
        end
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
    | _ => no_tac)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   308
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   309
(* Injection means to prove that the regularised theorem implies
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   310
   the abs/rep injected one.
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   312
   The deterministic part:
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   313
    - remove lambdas from both sides
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   314
    - prove Ball/Bex/Babs equalities using ball_rsp, bex_rsp, babs_rsp
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   315
    - prove Ball/Bex relations unfolding fun_rel_id
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   316
    - reflexivity of equality
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   317
    - prove equality of relations using equals_rsp
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   318
    - use user-supplied RSP theorems
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   319
    - solve 'relation of relations' goals using quot_rel_rsp
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   320
    - remove rep_abs from the right side
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   321
      (Lambdas under respects may have left us some assumptions)
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   322
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   323
   Then in order:
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   324
    - split applications of lifted type (apply_rsp)
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   325
    - split applications of non-lifted type (cong_tac)
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   326
    - apply extentionality
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   327
    - assumption
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   328
    - reflexivity of the relation
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
*)
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   330
fun injection_match_tac ctxt = SUBGOAL (fn (goal, i) =>
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
(case (bare_concl goal) of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   332
    (* (R1 ===> R2) (%x...) (%x...) ----> [|R1 x y|] ==> R2 (...x) (...y) *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
  (Const (@{const_name fun_rel}, _) $ _ $ _) $ (Abs _) $ (Abs _)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   334
      => rtac @{thm fun_rel_id} THEN' quot_true_tac ctxt unlam
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   336
    (* (op =) (Ball...) (Ball...) ----> (op =) (...) (...) *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
| (Const (@{const_name "op ="},_) $
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
    (Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
    (Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _))
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
      => rtac @{thm ball_rsp} THEN' dtac @{thm QT_all}
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   341
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   342
    (* (R1 ===> op =) (Ball...) (Ball...) ----> [|R1 x y|] ==> (Ball...x) = (Ball...y) *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   343
| (Const (@{const_name fun_rel}, _) $ _ $ _) $
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   344
    (Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   345
    (Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
      => rtac @{thm fun_rel_id} THEN' quot_true_tac ctxt unlam
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   347
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
    (* (op =) (Bex...) (Bex...) ----> (op =) (...) (...) *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   349
| Const (@{const_name "op ="},_) $
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   350
    (Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   351
    (Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   352
      => rtac @{thm bex_rsp} THEN' dtac @{thm QT_ex}
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   353
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   354
    (* (R1 ===> op =) (Bex...) (Bex...) ----> [|R1 x y|] ==> (Bex...x) = (Bex...y) *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   355
| (Const (@{const_name fun_rel}, _) $ _ $ _) $
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   356
    (Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   357
    (Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   358
      => rtac @{thm fun_rel_id} THEN' quot_true_tac ctxt unlam
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   359
911
95ee248b3832 Automatic injection of Bexeq
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 910
diff changeset
   360
| (Const (@{const_name fun_rel}, _) $ _ $ _) $
980
9d35c6145dd2 Renamed Bexeq to Bex1_rel
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 964
diff changeset
   361
    (Const(@{const_name Bex1_rel},_) $ _) $ (Const(@{const_name Bex1_rel},_) $ _)
1120
42b623872781 Finishe the renaming.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1113
diff changeset
   362
      => rtac @{thm bex1_rel_rsp} THEN' quotient_tac ctxt
911
95ee248b3832 Automatic injection of Bexeq
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 910
diff changeset
   363
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   364
| (_ $
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   365
    (Const(@{const_name Babs},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   366
    (Const(@{const_name Babs},_) $ (Const (@{const_name Respects}, _) $ _) $ _))
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   367
      => rtac @{thm babs_rsp} THEN' RANGE [quotient_tac ctxt]
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   368
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   369
| Const (@{const_name "op ="},_) $ (R $ _ $ _) $ (_ $ _ $ _) =>
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   370
   (rtac @{thm refl} ORELSE'
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   371
    (equals_rsp_tac R ctxt THEN' RANGE [
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   372
       quot_true_tac ctxt (fst o dest_bcomb), quot_true_tac ctxt (snd o dest_bcomb)]))
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   373
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   374
    (* reflexivity of operators arising from Cong_tac *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   375
| Const (@{const_name "op ="},_) $ _ $ _ => rtac @{thm refl}
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   376
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   377
   (* respectfulness of constants; in particular of a simple relation *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   378
| _ $ (Const _) $ (Const _)  (* fun_rel, list_rel, etc but not equality *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   379
    => resolve_tac (rsp_rules_get ctxt) THEN_ALL_NEW quotient_tac ctxt
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   380
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   381
    (* R (...) (Rep (Abs ...)) ----> R (...) (...) *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   382
    (* observe fun_map *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   383
| _ $ _ $ _
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   384
    => (rtac @{thm quot_rel_rsp} THEN_ALL_NEW quotient_tac ctxt)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   385
       ORELSE' rep_abs_rsp_tac ctxt
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   386
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   387
| _ => K no_tac
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   388
) i)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   389
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   390
fun injection_step_tac ctxt rel_refl =
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   391
 FIRST' [
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   392
    injection_match_tac ctxt,
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   393
853
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   394
    (* R (t $ ...) (t' $ ...) ----> apply_rsp   provided type of t needs lifting *)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   395
    apply_rsp_tac ctxt THEN'
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   396
                 RANGE [quot_true_tac ctxt (fst o dest_comb), quot_true_tac ctxt (snd o dest_comb)],
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   397
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   398
    (* (op =) (t $ ...) (t' $ ...) ----> Cong   provided type of t does not need lifting *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   399
    (* merge with previous tactic *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   400
    Cong_Tac.cong_tac @{thm cong} THEN'
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   401
                 RANGE [quot_true_tac ctxt (fst o dest_comb), quot_true_tac ctxt (snd o dest_comb)],
853
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   402
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   403
    (* (op =) (%x...) (%y...) ----> (op =) (...) (...) *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   404
    rtac @{thm ext} THEN' quot_true_tac ctxt unlam,
853
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   405
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   406
    (* resolving with R x y assumptions *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   407
    atac,
853
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   408
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   409
    (* reflexivity of the basic relations *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   410
    (* R ... ... *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   411
    resolve_tac rel_refl]
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   412
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   413
fun injection_tac ctxt =
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   414
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   415
  val rel_refl = map (OF1 @{thm equivp_reflp}) (equiv_rules_get ctxt)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   416
in
845
9531fafc0daa removed 3 hacks.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 844
diff changeset
   417
  injection_step_tac ctxt rel_refl
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   418
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   419
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   420
fun all_injection_tac ctxt =
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   421
  REPEAT_ALL_NEW (injection_tac ctxt)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   422
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   423
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   424
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   425
(*** Cleaning of the Theorem ***)
848
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   426
951
62f0344b219c some slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 946
diff changeset
   427
(* expands all fun_maps, except in front of the (bound) variables listed in xs *)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   428
fun fun_map_simple_conv xs ctrm =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   429
  case (term_of ctrm) of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   430
    ((Const (@{const_name "fun_map"}, _) $ _ $ _) $ h $ _) =>
1102
26f15c2dc131 removing unnecessary brackets
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1084
diff changeset
   431
        if member (op=) xs h
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   432
        then Conv.all_conv ctrm
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   433
        else Conv.rewr_conv @{thm fun_map_def[THEN eq_reflection]} ctrm
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   434
  | _ => Conv.all_conv ctrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   435
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   436
fun fun_map_conv xs ctxt ctrm =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   437
  case (term_of ctrm) of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   438
      _ $ _ => (Conv.comb_conv (fun_map_conv xs ctxt) then_conv
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   439
                fun_map_simple_conv xs) ctrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   440
    | Abs _ => Conv.abs_conv (fn (x, ctxt) => fun_map_conv ((term_of x)::xs) ctxt) ctxt ctrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   441
    | _ => Conv.all_conv ctrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   442
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   443
fun fun_map_tac ctxt = CONVERSION (fun_map_conv [] ctxt)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   444
951
62f0344b219c some slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 946
diff changeset
   445
(* custom matching functions *)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   446
fun mk_abs u i t =
1112
c7069b09730b More changes according to the standards.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1102
diff changeset
   447
  if incr_boundvars i u aconv t then Bound i else
c7069b09730b More changes according to the standards.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1102
diff changeset
   448
  case t of
c7069b09730b More changes according to the standards.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1102
diff changeset
   449
    t1 $ t2 => mk_abs u i t1 $ mk_abs u i t2
c7069b09730b More changes according to the standards.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1102
diff changeset
   450
  | Abs (s, T, t') => Abs (s, T, mk_abs u (i + 1) t')
c7069b09730b More changes according to the standards.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1102
diff changeset
   451
  | Bound j => if i = j then error "make_inst" else t
c7069b09730b More changes according to the standards.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1102
diff changeset
   452
  | _ => t
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   453
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   454
fun make_inst lhs t =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   455
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   456
  val _ $ (Abs (_, _, (_ $ ((f as Var (_, Type ("fun", [T, _]))) $ u)))) = lhs;
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   457
  val _ $ (Abs (_, _, (_ $ g))) = t;
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   458
in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   459
  (f, Abs ("x", T, mk_abs u 0 g))
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   460
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   461
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   462
fun make_inst_id lhs t =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   463
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   464
  val _ $ (Abs (_, _, (f as Var (_, Type ("fun", [T, _]))) $ u)) = lhs;
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   465
  val _ $ (Abs (_, _, g)) = t;
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   466
in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   467
  (f, Abs ("x", T, mk_abs u 0 g))
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   468
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   469
846
9a0a0b92e8fe handle all is no longer necessary in lambda_prs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 845
diff changeset
   470
(* Simplifies a redex using the 'lambda_prs' theorem.
9a0a0b92e8fe handle all is no longer necessary in lambda_prs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 845
diff changeset
   471
   First instantiates the types and known subterms.
9a0a0b92e8fe handle all is no longer necessary in lambda_prs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 845
diff changeset
   472
   Then solves the quotient assumptions to get Rep2 and Abs1
9a0a0b92e8fe handle all is no longer necessary in lambda_prs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 845
diff changeset
   473
   Finally instantiates the function f using make_inst
9a0a0b92e8fe handle all is no longer necessary in lambda_prs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 845
diff changeset
   474
   If Rep2 is an identity then the pattern is simpler and
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   475
   make_inst_id is used
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   476
*)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   477
fun lambda_prs_simple_conv ctxt ctrm =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   478
  case (term_of ctrm) of
846
9a0a0b92e8fe handle all is no longer necessary in lambda_prs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 845
diff changeset
   479
    (Const (@{const_name fun_map}, _) $ r1 $ a2) $ (Abs _) =>
9a0a0b92e8fe handle all is no longer necessary in lambda_prs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 845
diff changeset
   480
      let
9a0a0b92e8fe handle all is no longer necessary in lambda_prs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 845
diff changeset
   481
        val thy = ProofContext.theory_of ctxt
9a0a0b92e8fe handle all is no longer necessary in lambda_prs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 845
diff changeset
   482
        val (ty_b, ty_a) = dest_fun_type (fastype_of r1)
9a0a0b92e8fe handle all is no longer necessary in lambda_prs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 845
diff changeset
   483
        val (ty_c, ty_d) = dest_fun_type (fastype_of a2)
9a0a0b92e8fe handle all is no longer necessary in lambda_prs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 845
diff changeset
   484
        val tyinst = map (SOME o (ctyp_of thy)) [ty_a, ty_b, ty_c, ty_d]
9a0a0b92e8fe handle all is no longer necessary in lambda_prs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 845
diff changeset
   485
        val tinst = [NONE, NONE, SOME (cterm_of thy r1), NONE, SOME (cterm_of thy a2)]
951
62f0344b219c some slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 946
diff changeset
   486
        val thm1 = Drule.instantiate' tyinst tinst @{thm lambda_prs[THEN eq_reflection]}
62f0344b219c some slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 946
diff changeset
   487
        val thm2 = solve_quotient_assm ctxt (solve_quotient_assm ctxt thm1)
62f0344b219c some slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 946
diff changeset
   488
        val thm3 = MetaSimplifier.rewrite_rule @{thms id_apply[THEN eq_reflection]} thm2
846
9a0a0b92e8fe handle all is no longer necessary in lambda_prs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 845
diff changeset
   489
        val (insp, inst) =
9a0a0b92e8fe handle all is no longer necessary in lambda_prs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 845
diff changeset
   490
          if ty_c = ty_d
951
62f0344b219c some slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 946
diff changeset
   491
          then make_inst_id (term_of (Thm.lhs_of thm3)) (term_of ctrm)
62f0344b219c some slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 946
diff changeset
   492
          else make_inst (term_of (Thm.lhs_of thm3)) (term_of ctrm)
62f0344b219c some slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 946
diff changeset
   493
        val thm4 = Drule.instantiate ([], [(cterm_of thy insp, cterm_of thy inst)]) thm3
846
9a0a0b92e8fe handle all is no longer necessary in lambda_prs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 845
diff changeset
   494
      in
951
62f0344b219c some slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 946
diff changeset
   495
        Conv.rewr_conv thm4 ctrm
846
9a0a0b92e8fe handle all is no longer necessary in lambda_prs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 845
diff changeset
   496
      end
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   497
  | _ => Conv.all_conv ctrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   498
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   499
fun lambda_prs_conv ctxt = More_Conv.top_conv lambda_prs_simple_conv ctxt
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   500
fun lambda_prs_tac ctxt = CONVERSION (lambda_prs_conv ctxt)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   501
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   502
848
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   503
(* Cleaning consists of:
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   504
940
a792bfc1be2b Combined the simpsets in clean_tac and updated the comment. Now cleaning of splits does work.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 939
diff changeset
   505
  1. unfolding of ---> in front of everything, except
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   506
     bound variables (this prevents lambda_prs from
951
62f0344b219c some slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 946
diff changeset
   507
     becoming stuck)
848
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   508
940
a792bfc1be2b Combined the simpsets in clean_tac and updated the comment. Now cleaning of splits does work.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 939
diff changeset
   509
  2. simplification with lambda_prs
848
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   510
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   511
  3. simplification with:
951
62f0344b219c some slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 946
diff changeset
   512
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   513
      - Quotient_abs_rep Quotient_rel_rep
951
62f0344b219c some slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 946
diff changeset
   514
        babs_prs all_prs ex_prs ex1_prs
964
31907e0648ee tuned comment
Christian Urban <urbanc@in.tum.de>
parents: 952
diff changeset
   515
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   516
      - id_simps and preservation lemmas and
964
31907e0648ee tuned comment
Christian Urban <urbanc@in.tum.de>
parents: 952
diff changeset
   517
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   518
      - symmetric versions of the definitions
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   519
        (that is definitions of quotient constants
964
31907e0648ee tuned comment
Christian Urban <urbanc@in.tum.de>
parents: 952
diff changeset
   520
         are folded)
848
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   521
940
a792bfc1be2b Combined the simpsets in clean_tac and updated the comment. Now cleaning of splits does work.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 939
diff changeset
   522
  4. test for refl
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   523
*)
845
9531fafc0daa removed 3 hacks.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 844
diff changeset
   524
fun clean_tac lthy =
785
Christian Urban <urbanc@in.tum.de>
parents: 781
diff changeset
   525
let
1157
7763756b42cf Removed varifyT.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1137
diff changeset
   526
  val defs = map (symmetric o #def) (qconsts_dest lthy)
789
Christian Urban <urbanc@in.tum.de>
parents: 785
diff changeset
   527
  val prs = prs_rules_get lthy
Christian Urban <urbanc@in.tum.de>
parents: 785
diff changeset
   528
  val ids = id_simps_get lthy
951
62f0344b219c some slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 946
diff changeset
   529
  val thms = @{thms Quotient_abs_rep Quotient_rel_rep babs_prs all_prs ex_prs ex1_prs} @ ids @ prs @ defs
853
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   530
951
62f0344b219c some slight tuning
Christian Urban <urbanc@in.tum.de>
parents: 946
diff changeset
   531
  val ss = (mk_minimal_ss lthy) addsimps thms addSolver quotient_solver
785
Christian Urban <urbanc@in.tum.de>
parents: 781
diff changeset
   532
in
940
a792bfc1be2b Combined the simpsets in clean_tac and updated the comment. Now cleaning of splits does work.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 939
diff changeset
   533
  EVERY' [fun_map_tac lthy,
785
Christian Urban <urbanc@in.tum.de>
parents: 781
diff changeset
   534
          lambda_prs_tac lthy,
940
a792bfc1be2b Combined the simpsets in clean_tac and updated the comment. Now cleaning of splits does work.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 939
diff changeset
   535
          simp_tac ss,
853
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   536
          TRY o rtac refl]
785
Christian Urban <urbanc@in.tum.de>
parents: 781
diff changeset
   537
end
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   538
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   539
848
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   540
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   541
(** Tactic for Generalising Free Variables in a Goal **)
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   542
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   543
fun inst_spec ctrm =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   544
   Drule.instantiate' [SOME (ctyp_of_term ctrm)] [NONE, SOME ctrm] @{thm spec}
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   545
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   546
fun inst_spec_tac ctrms =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   547
  EVERY' (map (dtac o inst_spec) ctrms)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   548
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   549
fun all_list xs trm =
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   550
  fold (fn (x, T) => fn t' => HOLogic.mk_all (x, T, t')) xs trm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   551
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   552
fun apply_under_Trueprop f =
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   553
  HOLogic.dest_Trueprop #> f #> HOLogic.mk_Trueprop
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   554
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   555
fun gen_frees_tac ctxt =
853
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   556
  SUBGOAL (fn (concl, i) =>
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   557
    let
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   558
      val thy = ProofContext.theory_of ctxt
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   559
      val vrs = Term.add_frees concl []
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   560
      val cvrs = map (cterm_of thy o Free) vrs
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   561
      val concl' = apply_under_Trueprop (all_list vrs) concl
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   562
      val goal = Logic.mk_implies (concl', concl)
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   563
      val rule = Goal.prove ctxt [] [] goal
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   564
        (K (EVERY1 [inst_spec_tac (rev cvrs), atac]))
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   565
    in
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   566
      rtac rule i
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   567
    end)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   568
848
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   569
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   570
(** The General Shape of the Lifting Procedure **)
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   571
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   572
(* - A is the original raw theorem
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   573
   - B is the regularized theorem
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   574
   - C is the rep/abs injected version of B
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   575
   - D is the lifted theorem
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   576
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   577
   - 1st prem is the regularization step
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   578
   - 2nd prem is the rep/abs injection step
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   579
   - 3rd prem is the cleaning part
848
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   580
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   581
   the Quot_True premise in 2nd records the lifted theorem
858
Christian Urban <urbanc@in.tum.de>
parents: 857
diff changeset
   582
*)
853
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   583
val lifting_procedure_thm =
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   584
  @{lemma  "[|A;
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   585
              A --> B;
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   586
              Quot_True D ==> B = C;
3fd1365f5729 More indenting, bracket removing and comment restructuring.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 848
diff changeset
   587
              C = D|] ==> D"
773
d6acae26d027 tuned comments; renamed QUOT_TRUE to Quot_True; atomize_eqv seems to not be neccessary (has it been added to Isabelle)...it is now comented out and everything still works
Christian Urban <urbanc@in.tum.de>
parents: 772
diff changeset
   588
      by (simp add: Quot_True_def)}
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   589
1112
c7069b09730b More changes according to the standards.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1102
diff changeset
   590
fun lift_match_error ctxt msg rtrm qtrm =
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   591
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   592
  val rtrm_str = Syntax.string_of_term ctxt rtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   593
  val qtrm_str = Syntax.string_of_term ctxt qtrm
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   594
  val msg = cat_lines [enclose "[" "]" msg, "The quotient theorem", qtrm_str,
848
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   595
    "", "does not match with original theorem", rtrm_str]
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   596
in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   597
  error msg
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   598
end
848
0eb018699f46 Cleaning comments, indentation etc in quotient_tacs.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 847
diff changeset
   599
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   600
fun procedure_inst ctxt rtrm qtrm =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   601
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   602
  val thy = ProofContext.theory_of ctxt
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   603
  val rtrm' = HOLogic.dest_Trueprop rtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   604
  val qtrm' = HOLogic.dest_Trueprop qtrm
776
d1064fa29424 renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents: 775
diff changeset
   605
  val reg_goal = regularize_trm_chk ctxt (rtrm', qtrm')
1450
1ae5afcddcd4 another synchronisation
Christian Urban <urbanc@in.tum.de>
parents: 1438
diff changeset
   606
    handle (ERROR msg) => lift_match_error ctxt msg rtrm qtrm
776
d1064fa29424 renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents: 775
diff changeset
   607
  val inj_goal = inj_repabs_trm_chk ctxt (reg_goal, qtrm')
1450
1ae5afcddcd4 another synchronisation
Christian Urban <urbanc@in.tum.de>
parents: 1438
diff changeset
   608
    handle (ERROR msg) => lift_match_error ctxt msg rtrm qtrm
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   609
in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   610
  Drule.instantiate' []
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   611
    [SOME (cterm_of thy rtrm'),
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   612
     SOME (cterm_of thy reg_goal),
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   613
     NONE,
802
Christian Urban <urbanc@in.tum.de>
parents: 789
diff changeset
   614
     SOME (cterm_of thy inj_goal)] lifting_procedure_thm
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   615
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   616
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   617
(* the tactic leaves three subgoals to be proved *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   618
fun procedure_tac ctxt rthm =
1354
367f67311e6f updated to renamings in Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 1260
diff changeset
   619
  Object_Logic.full_atomize_tac
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   620
  THEN' gen_frees_tac ctxt
785
Christian Urban <urbanc@in.tum.de>
parents: 781
diff changeset
   621
  THEN' SUBGOAL (fn (goal, i) =>
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   622
    let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   623
      val rthm' = atomize_thm rthm
785
Christian Urban <urbanc@in.tum.de>
parents: 781
diff changeset
   624
      val rule = procedure_inst ctxt (prop_of rthm') goal
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   625
    in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   626
      (rtac rule THEN' rtac rthm') i
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   627
    end)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   628
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   629
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   630
(* Automatic Proofs *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   631
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 871
diff changeset
   632
val msg1 = "The regularize proof failed."
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   633
val msg2 = cat_lines ["The injection proof failed.",
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   634
                      "This is probably due to missing respects lemmas.",
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   635
                      "Try invoking the injection method manually to see",
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   636
                      "which lemmas are missing."]
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 871
diff changeset
   637
val msg3 = "The cleaning proof failed."
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   638
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 871
diff changeset
   639
fun lift_tac ctxt rthms =
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 871
diff changeset
   640
let
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   641
  fun mk_tac rthm =
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 871
diff changeset
   642
    procedure_tac ctxt rthm
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   643
    THEN' RANGE_WARN
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 871
diff changeset
   644
      [(regularize_tac ctxt, msg1),
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 871
diff changeset
   645
       (all_injection_tac ctxt, msg2),
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 871
diff changeset
   646
       (clean_tac ctxt, msg3)]
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 871
diff changeset
   647
in
1128
17ca92ab4660 Main renaming + fixes for new Isabelle in IntEx2.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1120
diff changeset
   648
  simp_tac (mk_minimal_ss ctxt) (* unfolding multiple &&& *)
896
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 871
diff changeset
   649
  THEN' RANGE (map mk_tac rthms)
4e0b89d886ac liftin and lifing_tac can now lift several "and"-separated goals at once; the raw-theorems have to be given in the order of goals
Christian Urban <urbanc@in.tum.de>
parents: 871
diff changeset
   650
end
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   651
1078
Christian Urban <urbanc@in.tum.de>
parents: 1077
diff changeset
   652
(* An Attribute which automatically constructs the qthm *)
Christian Urban <urbanc@in.tum.de>
parents: 1077
diff changeset
   653
fun lifted_attrib_aux context thm =
1068
62e54830590f A proper version of the attribute
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 980
diff changeset
   654
let
1077
44461d5615eb some tuning
Christian Urban <urbanc@in.tum.de>
parents: 1075
diff changeset
   655
  val ctxt = Context.proof_of context
1082
f8029d8c88a9 Fixed the context import/export and simplified LFex.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1078
diff changeset
   656
  val ((_, [thm']), ctxt') = Variable.import false [thm] ctxt
1083
30550327651a Proper context fixes lifting inside instantiations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1082
diff changeset
   657
  val goal = (quotient_lift_all ctxt' o prop_of) thm'
1068
62e54830590f A proper version of the attribute
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 980
diff changeset
   658
in
1084
40e3e6a6076f slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 1083
diff changeset
   659
  Goal.prove ctxt' [] [] goal (K (lift_tac ctxt' [thm] 1))
40e3e6a6076f slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 1083
diff changeset
   660
  |> singleton (ProofContext.export ctxt' ctxt)
1068
62e54830590f A proper version of the attribute
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 980
diff changeset
   661
end;
62e54830590f A proper version of the attribute
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 980
diff changeset
   662
1078
Christian Urban <urbanc@in.tum.de>
parents: 1077
diff changeset
   663
val lifted_attrib = Thm.rule_attribute lifted_attrib_aux
1068
62e54830590f A proper version of the attribute
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 980
diff changeset
   664
814
cd3fa86be45f Sledgehammer bug.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 804
diff changeset
   665
end; (* structure *)