Quot/quotient_tacs.ML
author Christian Urban <urbanc@in.tum.de>
Mon, 21 Dec 2009 22:36:31 +0100
changeset 768 e9e205b904e2
parent 762 baac4639ecef
child 769 d89851ebac9b
permissions -rw-r--r--
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
signature QUOTIENT_TACS =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
sig
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
    val regularize_tac: Proof.context -> int -> tactic
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
    val all_inj_repabs_tac: Proof.context -> int -> tactic
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
    val clean_tac: Proof.context -> int -> tactic
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
    val procedure_tac: Proof.context -> thm -> int -> tactic
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
    val lift_tac: Proof.context ->thm -> int -> tactic
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
    val quotient_tac: Proof.context -> int -> tactic
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
end;
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
structure Quotient_Tacs: QUOTIENT_TACS =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
struct
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
762
baac4639ecef avoided global "open"s - replaced by local "open"s
Christian Urban <urbanc@in.tum.de>
parents: 761
diff changeset
    14
open Quotient_Info;
baac4639ecef avoided global "open"s - replaced by local "open"s
Christian Urban <urbanc@in.tum.de>
parents: 761
diff changeset
    15
open Quotient_Type;
baac4639ecef avoided global "open"s - replaced by local "open"s
Christian Urban <urbanc@in.tum.de>
parents: 761
diff changeset
    16
open Quotient_Term;
baac4639ecef avoided global "open"s - replaced by local "open"s
Christian Urban <urbanc@in.tum.de>
parents: 761
diff changeset
    17
baac4639ecef avoided global "open"s - replaced by local "open"s
Christian Urban <urbanc@in.tum.de>
parents: 761
diff changeset
    18
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
(* Since HOL_basic_ss is too "big" for us,    *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
(* we need to set up our own minimal simpset. *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
fun  mk_minimal_ss ctxt =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  Simplifier.context ctxt empty_ss
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
    setsubgoaler asm_simp_tac
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
    setmksimps (mksimps [])
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
(* various helper functions *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
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
    30
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
(* makes sure a subgoal is solved *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
fun SOLVES' tac = tac THEN_ALL_NEW (K no_tac)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
(* prints warning, if goal is unsolved *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
fun WARN (tac, msg) i st =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
 case Seq.pull ((SOLVES' tac) i st) of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
     NONE    => (warning msg; Seq.single st)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
   | seqcell => Seq.make (fn () => seqcell)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
fun RANGE_WARN xs = RANGE (map WARN xs)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
fun atomize_thm thm =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
  val thm' = Thm.freezeT (forall_intr_vars thm)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
  val thm'' = ObjectLogic.atomize (cprop_of thm')
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
  @{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
    48
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
(* Regularize Tactic *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
fun equiv_tac ctxt =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
  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
    57
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
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
    59
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
    60
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
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
    62
  (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
    63
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
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
    65
  (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
    66
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
fun matching_prs thy pat trm =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
  val univ = Unify.matchers thy [(pat, trm)]
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
  val SOME (env, _) = Seq.pull univ
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
  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
    72
  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
    73
in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
  (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
    75
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
fun calculate_instance ctxt thm redex R1 R2 =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
  val thy = ProofContext.theory_of ctxt
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
  val goal = Const (@{const_name "equivp"}, dummyT) $ R2  
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
             |> Syntax.check_term ctxt
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
             |> HOLogic.mk_Trueprop 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
  val eqv_prem = Goal.prove ctxt [] [] goal (fn _ => equiv_tac ctxt 1)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
  val thm = (@{thm eq_reflection} OF [thm OF [eqv_prem]])
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
  val R1c = cterm_of thy R1
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
  val thmi = Drule.instantiate' [] [SOME R1c] thm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
  val inst = matching_prs thy (term_of (Thm.lhs_of thmi)) redex
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
  val thm2 = Drule.eta_contraction_rule (Drule.instantiate inst thmi)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
  SOME thm2
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
handle _ => NONE
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
(* FIXME/TODO: what is the place where the exception is raised: matching_prs? *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
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
    96
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
  val ctxt = Simplifier.the_context ss
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
in 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
 case redex of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
    (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
   101
      (Const (@{const_name "fun_rel"}, _) $ R1 $ R2)) $ _) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
        calculate_instance ctxt @{thm ball_reg_eqv_range} redex R1 R2
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
  | (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
   105
      (Const (@{const_name "fun_rel"}, _) $ R1 $ R2)) $ _) =>  
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
        calculate_instance ctxt @{thm bex_reg_eqv_range} redex R1 R2
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
  | _ => NONE
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
(* test whether DETERM makes any difference *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
fun quotient_tac ctxt = SOLVES'  
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
  (REPEAT_ALL_NEW (FIRST'
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
    [rtac @{thm identity_quotient},
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
     resolve_tac (quotient_rules_get ctxt)]))
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
fun quotient_solver_tac ss = quotient_tac (Simplifier.the_context ss)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
val quotient_solver = Simplifier.mk_solver' "Quotient goal solver" quotient_solver_tac
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
fun solve_quotient_assum ctxt thm =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
  case Seq.pull (quotient_tac ctxt 1 thm) of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
    SOME (t, _) => t
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
  | _ => error "solve_quotient_assum failed. Maybe a quotient_thm is missing"
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
(* 0. preliminary simplification step according to *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
(*    thm ball_reg_eqv bex_reg_eqv babs_reg_eqv    *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
(*        ball_reg_eqv_range bex_reg_eqv_range     *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
(*                                                 *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
(* 1. eliminating simple Ball/Bex instances        *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
(*    thm ball_reg_right bex_reg_left              *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
(*                                                 *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
(* 2. monos                                        *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
(* 3. commutation rules for ball and bex           *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
(*    thm ball_all_comm bex_ex_comm                *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
(*                                                 *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
(* 4. then rel-equality (which need to be          *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
(*    instantiated to avoid loops)                 *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
(*    thm eq_imp_rel                               *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
(*                                                 *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
(* 5. then simplification like 0                   *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
(*                                                 *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
(* finally jump back to 1                          *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
fun regularize_tac ctxt =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
  val thy = ProofContext.theory_of ctxt
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
  val pat_ball = @{term "Ball (Respects (R1 ===> R2)) P"}
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
  val pat_bex  = @{term "Bex (Respects (R1 ===> R2)) P"}
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
  val simproc = Simplifier.simproc_i thy "" [pat_ball, pat_bex] (K (ball_bex_range_simproc))
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
  val simpset = (mk_minimal_ss ctxt) 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
                       addsimps @{thms ball_reg_eqv bex_reg_eqv babs_reg_eqv babs_simp}
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
                       addsimprocs [simproc] addSolver equiv_solver addSolver quotient_solver
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
  val eq_eqvs = map (OF1 @{thm eq_imp_rel}) (equiv_rules_get ctxt)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
  simp_tac simpset THEN'
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
  REPEAT_ALL_NEW (CHANGED o FIRST' [
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
    resolve_tac @{thms ball_reg_right bex_reg_left},
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
    resolve_tac (Inductive.get_monos ctxt),
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
    resolve_tac @{thms ball_all_comm bex_ex_comm},
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
    resolve_tac eq_eqvs,  
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
    simp_tac simpset])
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
(* Injection Tactic *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
(* looks for QUOT_TRUE assumtions, and in case its parameter   *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
(* is an application, it returns the function and the argument *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
fun find_qt_asm asms =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
  fun find_fun trm =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
    case trm of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
      (Const(@{const_name Trueprop}, _) $ (Const (@{const_name QUOT_TRUE}, _) $ _)) => true
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
    | _ => false
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
 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
   178
   SOME (_ $ (_ $ (f $ a))) => SOME (f, a)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
 | _ => NONE
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
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
   183
  case (term_of ctrm) of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
    (Const (@{const_name QUOT_TRUE}, _) $ x) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
    let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
      val fx = fnctn x;
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
      val thy = ProofContext.theory_of ctxt;
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
      val cx = cterm_of thy x;
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
      val cfx = cterm_of thy fx;
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
      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
   191
      val cfxt = ctyp_of thy (fastype_of fx);
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
      val thm = Drule.instantiate' [SOME cxt, SOME cfxt] [SOME cx, SOME cfx] @{thm QUOT_TRUE_imp}
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
    in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
      Conv.rewr_conv thm ctrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
    end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
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
   198
  case (term_of ctrm) of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
    (Const (@{const_name QUOT_TRUE}, _) $ _) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
      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
   201
  | _ $ _ => 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
   202
  | 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
   203
  | _ => Conv.all_conv ctrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
fun quot_true_tac ctxt fnctn = 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
   CONVERSION
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
    ((Conv.params_conv ~1 (fn ctxt =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
       (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
   209
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
fun dest_comb (f $ a) = (f, a) 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
fun dest_bcomb ((_ $ l) $ r) = (l, r) 
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
(* TODO: Can this be done easier? *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
fun unlam t =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
  case t of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
    (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
   217
  | _ => 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
   218
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
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
   220
  | 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
   221
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
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
   223
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
(* we apply apply_rsp only in case if the type needs lifting,      *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
(* which is the case if the type of the data in the QUOT_TRUE      *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
(* assumption is different from the corresponding type in the goal *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
val apply_rsp_tac =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
  Subgoal.FOCUS (fn {concl, asms, context,...} =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
  let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
    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
   232
    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
   233
  in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
    case (bare_concl, qt_asm) of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
      (R2 $ (f $ x) $ (g $ y), SOME (qt_fun, qt_arg)) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
         if (fastype_of qt_fun) = (fastype_of f) 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
         then no_tac                             
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
         else                                    
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
           let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
             val ty_x = fastype_of x
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
             val ty_b = fastype_of qt_arg
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
             val ty_f = range_type (fastype_of f) 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
             val thy = ProofContext.theory_of context
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
             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
   245
             val t_inst = map (SOME o (cterm_of thy)) [R2, f, g, x, y];
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
             val inst_thm = Drule.instantiate' ty_inst ([NONE, NONE, NONE] @ t_inst) @{thm apply_rsp}
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
           in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
             (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
   249
           end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
    | _ => no_tac
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
  end)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
fun equals_rsp_tac R ctxt =
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 = domain_type (fastype_of R);
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
  val thy = ProofContext.theory_of ctxt
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
  val thm = Drule.instantiate' 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
               [SOME (ctyp_of thy ty)] [SOME (cterm_of thy R)] @{thm equals_rsp}
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
  rtac thm THEN' quotient_tac ctxt
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
(* raised by instantiate' *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
handle THM _  => K no_tac  
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
     | TYPE _ => K no_tac    
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
     | TERM _ => K no_tac
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
fun rep_abs_rsp_tac ctxt = 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   269
  SUBGOAL (fn (goal, i) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   270
    case (bare_concl goal) of 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
      (rel $ _ $ (rep $ (abs $ _))) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
        (let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
           val thy = ProofContext.theory_of ctxt;
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
           val (ty_a, ty_b) = dest_fun_type (fastype_of abs);
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
           val ty_inst = map (SOME o (ctyp_of thy)) [ty_a, ty_b];
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
           val t_inst = map (SOME o (cterm_of thy)) [rel, abs, rep];
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   277
           val inst_thm = Drule.instantiate' ty_inst t_inst @{thm rep_abs_rsp}
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
         in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
           (rtac inst_thm THEN' quotient_tac ctxt) i
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
         end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
         handle THM _ => no_tac | TYPE _ => no_tac)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
    | _ => no_tac)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
(* FIXME /TODO needs to be adapted *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   286
(*
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   287
To prove that the regularised theorem implies the abs/rep injected, 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
we try:
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   289
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
 1) theorems 'trans2' from the appropriate QUOT_TYPE
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   291
 2) remove lambdas from both sides: lambda_rsp_tac
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   292
 3) remove Ball/Bex from the right hand side
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   293
 4) use user-supplied RSP theorems
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   294
 5) remove rep_abs from the right side
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   295
 6) reflexivity of equality
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   296
 7) split applications of lifted type (apply_rsp)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   297
 8) split applications of non-lifted type (cong_tac)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   298
 9) apply extentionality
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   299
 A) reflexivity of the relation
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   300
 B) assumption
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   301
    (Lambdas under respects may have left us some assumptions)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   302
 C) proving obvious higher order equalities by simplifying fun_rel
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   303
    (not sure if it is still needed?)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   304
 D) unfolding lambda on one side
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
 E) simplifying (= ===> =) for simpler respectfulness
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
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   308
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   309
fun inj_repabs_tac_match ctxt = SUBGOAL (fn (goal, i) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   310
(case (bare_concl goal) of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
    (* (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
   312
  (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
   313
      => 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
   314
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   315
    (* (op =) (Ball...) (Ball...) ----> (op =) (...) (...) *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   316
| (Const (@{const_name "op ="},_) $
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   317
    (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
   318
    (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
   319
      => 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
   320
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
    (* (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
   322
| (Const (@{const_name fun_rel}, _) $ _ $ _) $
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   323
    (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
   324
    (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
   325
      => 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
   326
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
    (* (op =) (Bex...) (Bex...) ----> (op =) (...) (...) *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
| Const (@{const_name "op ="},_) $
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
    (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
   330
    (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
   331
      => 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
   332
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
    (* (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
   334
| (Const (@{const_name fun_rel}, _) $ _ $ _) $
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
    (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
   336
    (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
   337
      => 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
   338
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
| (_ $
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
    (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
   341
    (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
   342
      => 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
   343
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   344
| Const (@{const_name "op ="},_) $ (R $ _ $ _) $ (_ $ _ $ _) => 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   345
   (rtac @{thm refl} ORELSE'
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
    (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
   347
       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
   348
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   349
    (* 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
   350
| 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
   351
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   352
   (* 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
   353
| _ $ (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
   354
    => 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
   355
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   356
    (* R (...) (Rep (Abs ...)) ----> R (...) (...) *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   357
    (* observe fun_map *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   358
| _ $ _ $ _
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   359
    => (rtac @{thm quot_rel_rsp} THEN_ALL_NEW quotient_tac ctxt) 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   360
       ORELSE' rep_abs_rsp_tac ctxt
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   361
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
| _ => K no_tac
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   363
) i)
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
fun inj_repabs_step_tac ctxt rel_refl =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   366
 FIRST' [
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   367
    inj_repabs_tac_match ctxt,
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   368
    (* R (t $ ...) (t' $ ...) ----> apply_rsp   provided type of t needs lifting *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   369
    
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   370
    apply_rsp_tac ctxt THEN'
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   371
                 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
   372
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   373
    (* (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
   374
    (* merge with previous tactic *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   375
    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
   376
                 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
   377
    
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   378
    (* (op =) (%x...) (%y...) ----> (op =) (...) (...) *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   379
    rtac @{thm ext} THEN' quot_true_tac ctxt unlam,
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
    (* resolving with R x y assumptions *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   382
    atac,
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   383
    
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   384
    (* reflexivity of the basic relations *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   385
    (* R ... ... *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   386
    resolve_tac rel_refl]
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   387
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   388
fun inj_repabs_tac ctxt =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   389
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   390
  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
   391
in
768
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
   392
  simp_tac ((mk_minimal_ss ctxt) addsimps (id_simps_get ctxt)) (* HACK? *) 
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
   393
  THEN' inj_repabs_step_tac ctxt rel_refl
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   394
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   395
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   396
fun all_inj_repabs_tac ctxt =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   397
  REPEAT_ALL_NEW (inj_repabs_tac ctxt)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   398
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   399
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   400
(* Cleaning of the Theorem *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   401
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   402
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   403
(* expands all fun_maps, except in front of bound variables *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   404
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
   405
  case (term_of ctrm) of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   406
    ((Const (@{const_name "fun_map"}, _) $ _ $ _) $ h $ _) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   407
        if (member (op=) xs h) 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   408
        then Conv.all_conv ctrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   409
        else Conv.rewr_conv @{thm fun_map.simps[THEN eq_reflection]} ctrm 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   410
  | _ => Conv.all_conv ctrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   411
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   412
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
   413
  case (term_of ctrm) of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   414
      _ $ _ => (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
   415
                fun_map_simple_conv xs) ctrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   416
    | 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
   417
    | _ => Conv.all_conv ctrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   418
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   419
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
   420
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   421
fun mk_abs u i t =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   422
  if incr_boundvars i u aconv t then Bound i
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   423
  else (case t of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   424
     t1 $ t2 => (mk_abs u i t1) $ (mk_abs u i t2)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   425
   | Abs (s, T, t') => Abs (s, T, mk_abs u (i + 1) t')
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   426
   | Bound j => if i = j then error "make_inst" else t
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   427
   | _ => t)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   428
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   429
fun make_inst lhs t =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   430
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   431
  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
   432
  val _ $ (Abs (_, _, (_ $ g))) = t;
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   433
in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   434
  (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
   435
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   436
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   437
fun make_inst_id lhs t =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   438
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   439
  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
   440
  val _ $ (Abs (_, _, g)) = t;
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   441
in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   442
  (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
   443
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   444
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   445
(* Simplifies a redex using the 'lambda_prs' theorem.        *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   446
(* First instantiates the types and known subterms.          *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   447
(* Then solves the quotient assumptions to get Rep2 and Abs1 *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   448
(* Finally instantiates the function f using make_inst       *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   449
(* If Rep2 is identity then the pattern is simpler and       *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   450
(* make_inst_id is used                                      *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   451
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
   452
  case (term_of ctrm) of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   453
   (Const (@{const_name fun_map}, _) $ r1 $ a2) $ (Abs _) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   454
     (let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   455
       val thy = ProofContext.theory_of ctxt
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   456
       val (ty_b, ty_a) = dest_fun_type (fastype_of r1)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   457
       val (ty_c, ty_d) = dest_fun_type (fastype_of a2)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   458
       val tyinst = map (SOME o (ctyp_of thy)) [ty_a, ty_b, ty_c, ty_d]
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   459
       val tinst = [NONE, NONE, SOME (cterm_of thy r1), NONE, SOME (cterm_of thy a2)]
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   460
       val lpi = Drule.instantiate' tyinst tinst @{thm lambda_prs}
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   461
       val te = @{thm eq_reflection} OF [solve_quotient_assum ctxt (solve_quotient_assum ctxt lpi)]
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   462
       val ts = MetaSimplifier.rewrite_rule (id_simps_get ctxt) te
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   463
       val make_inst = if ty_c = ty_d then make_inst_id else make_inst
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   464
       val (insp, inst) = make_inst (term_of (Thm.lhs_of ts)) (term_of ctrm)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   465
       val ti = Drule.instantiate ([], [(cterm_of thy insp, cterm_of thy inst)]) ts
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
       Conv.rewr_conv ti ctrm
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
     handle _ => Conv.all_conv ctrm)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   470
  | _ => Conv.all_conv ctrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   471
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   472
val lambda_prs_conv =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   473
  More_Conv.top_conv lambda_prs_simple_conv
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   474
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   475
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
   476
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   477
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   478
(* 1. folding of definitions and preservation lemmas;  *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   479
(*    and simplification with                          *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   480
(*    thm babs_prs all_prs ex_prs                      *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   481
(*                                                     *) 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   482
(* 2. unfolding of ---> in front of everything, except *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   483
(*    bound variables (this prevents lambda_prs from   *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   484
(*    becoming stuck                                   *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   485
(*    thm fun_map.simps                                *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   486
(*                                                     *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   487
(* 3. simplification with                              *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   488
(*    thm lambda_prs                                   *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   489
(*                                                     *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   490
(* 4. simplification with                              *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   491
(*    thm Quotient_abs_rep Quotient_rel_rep id_simps   *) 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   492
(*                                                     *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   493
(* 5. Test for refl                                    *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   494
768
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
   495
fun clean_tac_aux lthy =
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   496
  let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   497
    val thy = ProofContext.theory_of lthy;
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   498
    val defs = map (Thm.varifyT o symmetric o #def) (qconsts_dest thy)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   499
      (* FIXME: why is the Thm.varifyT needed: example where it fails is LamEx *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   500
    
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   501
    val thms1 = defs @ (prs_rules_get lthy) @ @{thms babs_prs all_prs ex_prs}
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   502
    val thms2 = @{thms Quotient_abs_rep Quotient_rel_rep} @ (id_simps_get lthy) 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   503
    fun simps thms = (mk_minimal_ss lthy) addsimps thms addSolver quotient_solver
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   504
  in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   505
    EVERY' [simp_tac (simps thms1),
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   506
            fun_map_tac lthy,
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   507
            lambda_prs_tac lthy,
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   508
            simp_tac (simps thms2),
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   509
            TRY o rtac refl]
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   510
  end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   511
768
e9e205b904e2 get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
   512
fun clean_tac lthy = REPEAT o CHANGED o (clean_tac_aux lthy) (* HACK?? *)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   513
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   514
(* Tactic for Genralisation of Free Variables in a Goal *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   515
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   516
fun inst_spec ctrm =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   517
   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
   518
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   519
fun inst_spec_tac ctrms =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   520
  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
   521
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   522
fun all_list xs trm = 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   523
  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
   524
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   525
fun apply_under_Trueprop f = 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   526
  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
   527
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   528
fun gen_frees_tac ctxt =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   529
 SUBGOAL (fn (concl, i) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   530
  let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   531
    val thy = ProofContext.theory_of ctxt
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   532
    val vrs = Term.add_frees concl []
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   533
    val cvrs = map (cterm_of thy o Free) vrs
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   534
    val concl' = apply_under_Trueprop (all_list vrs) concl
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   535
    val goal = Logic.mk_implies (concl', concl)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   536
    val rule = Goal.prove ctxt [] [] goal 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   537
      (K (EVERY1 [inst_spec_tac (rev cvrs), atac]))
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   538
  in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   539
    rtac rule i
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   540
  end)  
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   541
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   542
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   543
(* The General Shape of the Lifting Procedure *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   544
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   545
(* - A is the original raw theorem                       *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   546
(* - B is the regularized theorem                        *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   547
(* - C is the rep/abs injected version of B              *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   548
(* - D is the lifted theorem                             *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   549
(*                                                       *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   550
(* - 1st prem is the regularization step                 *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   551
(* - 2nd prem is the rep/abs injection step              *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   552
(* - 3rd prem is the cleaning part                       *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   553
(*                                                       *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   554
(* the QUOT_TRUE premise in 2 records the lifted theorem *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   555
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   556
val lifting_procedure = 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   557
   @{lemma  "[|A; 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   558
               A --> B; 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   559
               QUOT_TRUE D ==> B = C; 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   560
               C = D|] ==> D" 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   561
      by (simp add: QUOT_TRUE_def)}
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   562
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   563
fun lift_match_error ctxt fun_str rtrm qtrm =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   564
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   565
  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
   566
  val qtrm_str = Syntax.string_of_term ctxt qtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   567
  val msg = cat_lines [enclose "[" "]" fun_str, "The quotient theorem", qtrm_str, 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   568
             "", "does not match with original theorem", rtrm_str]
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   569
in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   570
  error msg
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   571
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   572
 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   573
fun procedure_inst ctxt rtrm qtrm =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   574
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   575
  val thy = ProofContext.theory_of ctxt
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   576
  val rtrm' = HOLogic.dest_Trueprop rtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   577
  val qtrm' = HOLogic.dest_Trueprop qtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   578
  val reg_goal = 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   579
        Syntax.check_term ctxt (regularize_trm ctxt rtrm' qtrm')
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   580
        handle (LIFT_MATCH s) => lift_match_error ctxt s rtrm qtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   581
  val inj_goal = 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   582
        Syntax.check_term ctxt (inj_repabs_trm ctxt (reg_goal, qtrm'))
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   583
        handle (LIFT_MATCH s) => lift_match_error ctxt s rtrm qtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   584
in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   585
  Drule.instantiate' []
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   586
    [SOME (cterm_of thy rtrm'),
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   587
     SOME (cterm_of thy reg_goal),
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   588
     NONE,
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   589
     SOME (cterm_of thy inj_goal)] lifting_procedure
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   590
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   591
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   592
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   593
(* 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
   594
fun procedure_tac ctxt rthm =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   595
  ObjectLogic.full_atomize_tac
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   596
  THEN' gen_frees_tac ctxt
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   597
  THEN' CSUBGOAL (fn (goal, i) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   598
    let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   599
      val rthm' = atomize_thm rthm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   600
      val rule = procedure_inst ctxt (prop_of rthm') (term_of goal)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   601
    in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   602
      (rtac rule THEN' rtac rthm') i
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   603
    end)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   604
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   605
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   606
(* Automatic Proofs *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   607
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   608
val msg1 = "Regularize proof failed."
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   609
val msg2 = cat_lines ["Injection proof failed.", 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   610
                      "This is probably due to missing respects lemmas.",
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   611
                      "Try invoking the injection method manually to see", 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   612
                      "which lemmas are missing."]
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   613
val msg3 = "Cleaning proof failed."
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   614
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   615
fun lift_tac ctxt rthm =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   616
  procedure_tac ctxt rthm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   617
  THEN' RANGE_WARN 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   618
     [(regularize_tac ctxt, msg1),
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   619
      (all_inj_repabs_tac ctxt, msg2),
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   620
      (clean_tac ctxt, msg3)]
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   621
762
baac4639ecef avoided global "open"s - replaced by local "open"s
Christian Urban <urbanc@in.tum.de>
parents: 761
diff changeset
   622
end; (* structure *)