src/HOL/Tools/Qelim/qelim.ML
author nipkow
Thu, 26 Jun 2025 17:30:33 +0200
changeset 82772 59b937edcff8
parent 80693 e451d6230535
permissions -rw-r--r--
merged
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
80692
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
     9
  val gen_qelim_conv: Proof.context -> conv -> conv -> conv -> (cterm -> 'a -> 'a) -> 'a ->
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    10
    ('a -> conv) -> ('a -> conv) -> ('a -> conv) -> conv
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    11
  val standard_qelim_conv: Proof.context ->
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    12
    (cterm list -> conv) -> (cterm list -> conv) ->
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    13
    (cterm list -> conv) -> conv
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    14
end;
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    15
23489
6c2156c478f6 made type conv pervasive;
wenzelm
parents: 23466
diff changeset
    16
structure Qelim: QELIM =
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    17
struct
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    18
61075
f6b0d827240e tuned -- avoid slightly odd @{cpat};
wenzelm
parents: 60818
diff changeset
    19
fun gen_qelim_conv ctxt precv postcv simpex_conv ins env atcv ncv qcv =
80692
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    20
  let
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    21
    fun conv env p =
80693
e451d6230535 tuned: more antiquotations;
wenzelm
parents: 80692
diff changeset
    22
      (case Thm.term_of p of
e451d6230535 tuned: more antiquotations;
wenzelm
parents: 80692
diff changeset
    23
        \<^Const_>\<open>conj for _ _\<close> => Conv.binop_conv (conv env) p
e451d6230535 tuned: more antiquotations;
wenzelm
parents: 80692
diff changeset
    24
      | \<^Const_>\<open>disj for _ _\<close> => Conv.binop_conv (conv env) p
e451d6230535 tuned: more antiquotations;
wenzelm
parents: 80692
diff changeset
    25
      | \<^Const_>\<open>implies for _ _\<close> => Conv.binop_conv (conv env) p
e451d6230535 tuned: more antiquotations;
wenzelm
parents: 80692
diff changeset
    26
      | \<^Const_>\<open>HOL.eq _ for _ _\<close> => Conv.binop_conv (conv env) p
e451d6230535 tuned: more antiquotations;
wenzelm
parents: 80692
diff changeset
    27
      | \<^Const_>\<open>Not for _\<close> => Conv.arg_conv (conv env) p
e451d6230535 tuned: more antiquotations;
wenzelm
parents: 80692
diff changeset
    28
      | \<^Const_>\<open>Ex _ for \<open>Abs (s, _, _)\<close>\<close> =>
80692
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    29
          let
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    30
            val (e,p0) = Thm.dest_comb p
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    31
            val (x,p') = Thm.dest_abs_global p0
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    32
            val env' = ins x env
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    33
            val th =
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    34
              Thm.abstract_rule s x ((conv env' then_conv ncv env') p')
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    35
              |> Drule.arg_cong_rule e
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    36
            val th' = simpex_conv (Thm.rhs_of th)
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    37
            val (_, r) = Thm.dest_equals (Thm.cprop_of th')
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    38
          in
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    39
            if Thm.is_reflexive th' then Thm.transitive th (qcv env (Thm.rhs_of th))
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    40
            else Thm.transitive (Thm.transitive th th') (conv env r)
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    41
          end
80693
e451d6230535 tuned: more antiquotations;
wenzelm
parents: 80692
diff changeset
    42
      | \<^Const_>\<open>Ex _ for _\<close> => (Thm.eta_long_conversion then_conv conv env) p
e451d6230535 tuned: more antiquotations;
wenzelm
parents: 80692
diff changeset
    43
      | \<^Const_>\<open>All _ for _\<close> =>
80692
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    44
          let
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    45
            val allT = Thm.ctyp_of_cterm (Thm.dest_fun p)
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    46
            val T = Thm.dest_ctyp0 (Thm.dest_ctyp0 allT)
80693
e451d6230535 tuned: more antiquotations;
wenzelm
parents: 80692
diff changeset
    47
            val P = Thm.dest_arg p
e451d6230535 tuned: more antiquotations;
wenzelm
parents: 80692
diff changeset
    48
            val th = \<^instantiate>\<open>'a = T and P in lemma "\<forall>x::'a. P x \<equiv> \<nexists>x. \<not> P x" by simp\<close>
80692
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    49
          in Thm.transitive th (conv env (Thm.rhs_of th)) end
80693
e451d6230535 tuned: more antiquotations;
wenzelm
parents: 80692
diff changeset
    50
      | _ => atcv env p)
80692
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    51
  in precv then_conv (conv env) then_conv postcv end
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    52
23489
6c2156c478f6 made type conv pervasive;
wenzelm
parents: 23466
diff changeset
    53
23904
chaieb
parents: 23855
diff changeset
    54
(* Instantiation of some parameter for most common cases *)
23489
6c2156c478f6 made type conv pervasive;
wenzelm
parents: 23466
diff changeset
    55
23904
chaieb
parents: 23855
diff changeset
    56
local
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45654
diff changeset
    57
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45654
diff changeset
    58
val ss =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45654
diff changeset
    59
  simpset_of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 61075
diff changeset
    60
   (put_simpset HOL_basic_ss \<^context>
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45654
diff changeset
    61
    addsimps @{thms simp_thms ex_simps all_simps all_not_ex not_all ex_disj_distrib});
80692
f65b65814b81 tuned whitespace;
wenzelm
parents: 74525
diff changeset
    62
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45654
diff changeset
    63
fun pcv ctxt = Simplifier.rewrite (put_simpset ss ctxt)
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    64
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45654
diff changeset
    65
in
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45654
diff changeset
    66
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45654
diff changeset
    67
fun standard_qelim_conv ctxt atcv ncv qcv p =
74269
f084d599bb44 clarified set of items with order of addition;
wenzelm
parents: 74249
diff changeset
    68
  let
f084d599bb44 clarified set of items with order of addition;
wenzelm
parents: 74249
diff changeset
    69
    val pcv = pcv ctxt
74274
36f2c4a5c21b clarified signature;
wenzelm
parents: 74269
diff changeset
    70
    val env = Cterms.list_set_rev (Cterms.build (Drule.add_frees_cterm p))
74269
f084d599bb44 clarified set of items with order of addition;
wenzelm
parents: 74249
diff changeset
    71
  in gen_qelim_conv ctxt pcv pcv pcv cons env atcv ncv qcv p end
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45654
diff changeset
    72
23904
chaieb
parents: 23855
diff changeset
    73
end;
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    74
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    75
end;