src/HOL/Tools/Qelim/qelim.ML
author haftmann
Sun, 22 Mar 2009 20:46:12 +0100
changeset 30657 db260dfd2d8c
parent 23904 ba6c806590f8
child 32603 e08fdd615333
permissions -rw-r--r--
more antiquotations
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23489
6c2156c478f6 made type conv pervasive;
wenzelm
parents: 23466
diff changeset
     1
(*  Title:      HOL/Tools/Qelim/qelim.ML
23904
chaieb
parents: 23855
diff changeset
     2
    Author:     Amine Chaieb, TU Muenchen
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
     3
23904
chaieb
parents: 23855
diff changeset
     4
Generic quantifier elimination conversions for HOL formulae.
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
     5
*)
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
     6
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
     7
signature QELIM =
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
     8
sig
23904
chaieb
parents: 23855
diff changeset
     9
 val gen_qelim_conv: conv -> conv -> conv -> (cterm -> 'a -> 'a) -> 'a 
chaieb
parents: 23855
diff changeset
    10
                     -> ('a -> conv) -> ('a -> conv) -> ('a -> conv) -> conv
chaieb
parents: 23855
diff changeset
    11
 val standard_qelim_conv: (cterm list -> conv) -> (cterm list -> conv) 
chaieb
parents: 23855
diff changeset
    12
                          -> (cterm list -> conv) -> conv
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    13
end;
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    14
23489
6c2156c478f6 made type conv pervasive;
wenzelm
parents: 23466
diff changeset
    15
structure Qelim: QELIM =
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    16
struct
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    17
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    18
open Conv;
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    19
23489
6c2156c478f6 made type conv pervasive;
wenzelm
parents: 23466
diff changeset
    20
val all_not_ex = mk_meta_eq @{thm "all_not_ex"};
6c2156c478f6 made type conv pervasive;
wenzelm
parents: 23466
diff changeset
    21
23524
123a45589e0a Generic QE need no Context anymore
chaieb
parents: 23489
diff changeset
    22
fun gen_qelim_conv precv postcv simpex_conv ins env atcv ncv qcv =
23489
6c2156c478f6 made type conv pervasive;
wenzelm
parents: 23466
diff changeset
    23
 let
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    24
  fun conv env p =
23489
6c2156c478f6 made type conv pervasive;
wenzelm
parents: 23466
diff changeset
    25
   case (term_of p) of
23904
chaieb
parents: 23855
diff changeset
    26
    Const(s,T)$_$_ => 
chaieb
parents: 23855
diff changeset
    27
       if domain_type T = HOLogic.boolT
30657
db260dfd2d8c more antiquotations
haftmann
parents: 23904
diff changeset
    28
          andalso member (op =) [@{const_name "op &"}, @{const_name "op |"},
db260dfd2d8c more antiquotations
haftmann
parents: 23904
diff changeset
    29
            @{const_name "op -->"}, @{const_name "op ="}] s
23904
chaieb
parents: 23855
diff changeset
    30
       then binop_conv (conv env) p 
chaieb
parents: 23855
diff changeset
    31
       else atcv env p
30657
db260dfd2d8c more antiquotations
haftmann
parents: 23904
diff changeset
    32
  | Const(@{const_name "Not"},_)$_ => arg_conv (conv env) p
db260dfd2d8c more antiquotations
haftmann
parents: 23904
diff changeset
    33
  | Const(@{const_name "Ex"},_)$Abs(s,_,_) =>
23489
6c2156c478f6 made type conv pervasive;
wenzelm
parents: 23466
diff changeset
    34
    let
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    35
     val (e,p0) = Thm.dest_comb p
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    36
     val (x,p') = Thm.dest_abs (SOME s) p0
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    37
     val env' = ins x env
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    38
     val th = Thm.abstract_rule s x ((conv env' then_conv ncv env') p')
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    39
                   |> Drule.arg_cong_rule e
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    40
     val th' = simpex_conv (Thm.rhs_of th)
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    41
     val (l,r) = Thm.dest_equals (cprop_of th')
23593
fd12f7d56bd7 renamed Conv.is_refl to Thm.is_reflexive;
wenzelm
parents: 23524
diff changeset
    42
    in if Thm.is_reflexive th' then Thm.transitive th (qcv env (Thm.rhs_of th))
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    43
       else Thm.transitive (Thm.transitive th th') (conv env r) end
30657
db260dfd2d8c more antiquotations
haftmann
parents: 23904
diff changeset
    44
  | Const(@{const_name "Ex"},_)$ _ => (Thm.eta_long_conversion then_conv conv env) p
db260dfd2d8c more antiquotations
haftmann
parents: 23904
diff changeset
    45
  | Const(@{const_name "All"},_)$_ =>
23489
6c2156c478f6 made type conv pervasive;
wenzelm
parents: 23466
diff changeset
    46
    let
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    47
     val p = Thm.dest_arg p
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    48
     val ([(_,T)],[]) = Thm.match (@{cpat "All"}, Thm.dest_fun p)
23489
6c2156c478f6 made type conv pervasive;
wenzelm
parents: 23466
diff changeset
    49
     val th = instantiate' [SOME T] [SOME p] all_not_ex
6c2156c478f6 made type conv pervasive;
wenzelm
parents: 23466
diff changeset
    50
    in transitive th (conv env (Thm.rhs_of th))
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    51
    end
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    52
  | _ => atcv env p
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    53
 in precv then_conv (conv env) then_conv postcv end
23489
6c2156c478f6 made type conv pervasive;
wenzelm
parents: 23466
diff changeset
    54
23904
chaieb
parents: 23855
diff changeset
    55
(* Instantiation of some parameter for most common cases *)
23489
6c2156c478f6 made type conv pervasive;
wenzelm
parents: 23466
diff changeset
    56
23904
chaieb
parents: 23855
diff changeset
    57
local
23489
6c2156c478f6 made type conv pervasive;
wenzelm
parents: 23466
diff changeset
    58
val pcv = Simplifier.rewrite
6c2156c478f6 made type conv pervasive;
wenzelm
parents: 23466
diff changeset
    59
                 (HOL_basic_ss addsimps (simp_thms @ ex_simps @ all_simps)
6c2156c478f6 made type conv pervasive;
wenzelm
parents: 23466
diff changeset
    60
                     @ [@{thm "all_not_ex"}, not_all, ex_disj_distrib]);
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    61
23904
chaieb
parents: 23855
diff changeset
    62
in fun standard_qelim_conv atcv ncv qcv p =
chaieb
parents: 23855
diff changeset
    63
      gen_qelim_conv pcv pcv pcv cons (Thm.add_cterm_frees p []) atcv ncv qcv p
chaieb
parents: 23855
diff changeset
    64
end;
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    65
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    66
end;