src/HOL/Decision_Procs/ferrack_tac.ML
author wenzelm
Mon, 10 Nov 2014 21:49:48 +0100
changeset 58963 26bf09b95dda
parent 58956 a816aa3ff391
child 59582 0fbed69ff081
permissions -rw-r--r--
proper context for assume_tac (atac remains as fall-back without context);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30439
57c68b3af2ea Updated paths in Decision_Procs comments and NEWS
hoelzl
parents: 30242
diff changeset
     1
(*  Title:      HOL/Decision_Procs/ferrack_tac.ML
29788
1b80ebe713a4 established session HOL-Reflection
haftmann
parents: 29265
diff changeset
     2
    Author:     Amine Chaieb, TU Muenchen
1b80ebe713a4 established session HOL-Reflection
haftmann
parents: 29265
diff changeset
     3
*)
1b80ebe713a4 established session HOL-Reflection
haftmann
parents: 29265
diff changeset
     4
31302
12677a808d43 proper signature constraint;
wenzelm
parents: 30510
diff changeset
     5
signature FERRACK_TAC =
12677a808d43 proper signature constraint;
wenzelm
parents: 30510
diff changeset
     6
sig
12677a808d43 proper signature constraint;
wenzelm
parents: 30510
diff changeset
     7
  val linr_tac: Proof.context -> bool -> int -> tactic
12677a808d43 proper signature constraint;
wenzelm
parents: 30510
diff changeset
     8
end
12677a808d43 proper signature constraint;
wenzelm
parents: 30510
diff changeset
     9
55498
cf829d10d1d4 proper signature;
wenzelm
parents: 54742
diff changeset
    10
structure Ferrack_Tac: FERRACK_TAC =
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    11
struct
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    12
26075
815f3ccc0b45 added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
huffman
parents: 23590
diff changeset
    13
val ferrack_ss = let val ths = [@{thm real_of_int_inject}, @{thm real_of_int_less_iff}, 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    14
                                @{thm real_of_int_le_iff}]
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    15
             in @{context} delsimps ths addsimps (map (fn th => th RS sym) ths)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    16
             end |> simpset_of;
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    17
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45654
diff changeset
    18
val binarith = @{thms arith_simps}
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45654
diff changeset
    19
val comp_arith = binarith @ @{thms simp_thms}
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    20
55498
cf829d10d1d4 proper signature;
wenzelm
parents: 54742
diff changeset
    21
fun prepare_for_linr q fm = 
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    22
  let
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    23
    val ps = Logic.strip_params fm
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    24
    val hs = map HOLogic.dest_Trueprop (Logic.strip_assums_hyp fm)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    25
    val c = HOLogic.dest_Trueprop (Logic.strip_assums_concl fm)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    26
    fun mk_all ((s, T), (P,n)) =
42083
e1209fc7ecdc added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
wenzelm
parents: 38558
diff changeset
    27
      if Term.is_dependent P then
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    28
        (HOLogic.all_const T $ Abs (s, T, P), n)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    29
      else (incr_boundvars ~1 P, n-1)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    30
    fun mk_all2 (v, t) = HOLogic.all_const (fastype_of v) $ lambda v t;
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    31
      val rhs = hs
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    32
(*    val (rhs,irhs) = List.partition (relevant (rev ps)) hs *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    33
    val np = length ps
33004
715566791eb0 always qualify NJ's old List.foldl/foldr in Isabelle/ML;
wenzelm
parents: 32960
diff changeset
    34
    val (fm',np) =  List.foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n)))
715566791eb0 always qualify NJ's old List.foldl/foldr in Isabelle/ML;
wenzelm
parents: 32960
diff changeset
    35
      (List.foldr HOLogic.mk_imp c rhs, np) ps
27436
9581777503e9 code antiquotation roaring ahead
haftmann
parents: 26939
diff changeset
    36
    val (vs, _) = List.partition (fn t => q orelse (type_of t) = HOLogic.natT)
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    37
      (Misc_Legacy.term_frees fm' @ Misc_Legacy.term_vars fm');
33004
715566791eb0 always qualify NJ's old List.foldl/foldr in Isabelle/ML;
wenzelm
parents: 32960
diff changeset
    38
    val fm2 = List.foldr mk_all2 fm' vs
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    39
  in (fm2, np + length vs, length rhs) end;
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    40
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    41
(*Object quantifier to meta --*)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    42
fun spec_step n th = if (n=0) then th else (spec_step (n-1) th) RS spec ;
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    43
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    44
(* object implication to meta---*)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    45
fun mp_step n th = if (n=0) then th else (mp_step (n-1) th) RS mp;
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    46
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    47
35233
6af1caf7be69 local Simplifier.context;
wenzelm
parents: 35232
diff changeset
    48
fun linr_tac ctxt q =
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 52131
diff changeset
    49
    Object_Logic.atomize_prems_tac ctxt
58956
a816aa3ff391 proper context for compose_tac, Splitter.split_tac (relevant for unify trace options);
wenzelm
parents: 56245
diff changeset
    50
        THEN' (REPEAT_DETERM o split_tac ctxt [@{thm split_min}, @{thm split_max}, @{thm abs_split}])
35233
6af1caf7be69 local Simplifier.context;
wenzelm
parents: 35232
diff changeset
    51
        THEN' SUBGOAL (fn (g, i) =>
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    52
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42083
diff changeset
    53
    val thy = Proof_Context.theory_of ctxt
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    54
    (* Transform the term*)
55498
cf829d10d1d4 proper signature;
wenzelm
parents: 54742
diff changeset
    55
    val (t,np,nh) = prepare_for_linr q g
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    56
    (* Some simpsets for dealing with mod div abs and nat*)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    57
    val simpset0 = put_simpset HOL_basic_ss ctxt addsimps comp_arith
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    58
    val ct = cterm_of thy (HOLogic.mk_Trueprop t)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    59
    (* Theorem for the nat --> int transformation *)
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    60
   val pre_thm = Seq.hd (EVERY
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    61
      [simp_tac simpset0 1,
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    62
       TRY (simp_tac (put_simpset ferrack_ss ctxt) 1)]
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36853
diff changeset
    63
      (Thm.trivial ct))
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58956
diff changeset
    64
    fun assm_tac i = REPEAT_DETERM_N nh (assume_tac ctxt i)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    65
    (* The result of the quantifier elimination *)
35233
6af1caf7be69 local Simplifier.context;
wenzelm
parents: 35232
diff changeset
    66
    val (th, tac) = case prop_of pre_thm of
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55498
diff changeset
    67
        Const (@{const_name Pure.imp}, _) $ (Const (@{const_name Trueprop}, _) $ t1) $ _ =>
52131
366fa32ee2a3 tuned signature;
wenzelm
parents: 51717
diff changeset
    68
    let val pth = linr_oracle (ctxt, Envir.eta_long [] t1)
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    69
    in 
55498
cf829d10d1d4 proper signature;
wenzelm
parents: 54742
diff changeset
    70
       ((pth RS iffD2) RS pre_thm,
cf829d10d1d4 proper signature;
wenzelm
parents: 54742
diff changeset
    71
        assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI i))
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    72
    end
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    73
      | _ => (pre_thm, assm_tac i)
35233
6af1caf7be69 local Simplifier.context;
wenzelm
parents: 35232
diff changeset
    74
  in rtac ((mp_step nh o spec_step np) th) i THEN tac end);
23264
324622260d29 Added twe Examples for Quantifier elimination ofer linear real arithmetic and over the mixed theory of linear real artihmetic with integers
chaieb
parents:
diff changeset
    75
23469
3f309f885d0b add thm antiquotations
huffman
parents: 23318
diff changeset
    76
end