src/Provers/quantifier1.ML
author blanchet
Tue, 17 Aug 2010 13:10:58 +0200
changeset 38457 b8760b6e7c65
parent 38452 abc655166d61
parent 38456 6769ccd90ad6
child 42361 23f352990944
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
35762
af3ff2ba4c54 removed old CVS Ids;
wenzelm
parents: 31197
diff changeset
     1
(*  Title:      Provers/quantifier1.ML
4319
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
     2
    Author:     Tobias Nipkow
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
     3
    Copyright   1997  TU Munich
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
     4
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
     5
Simplification procedures for turning
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
     6
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
     7
            ? x. ... & x = t & ...
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
     8
     into   ? x. x = t & ... & ...
11232
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 11221
diff changeset
     9
     where the `? x. x = t &' in the latter formula must be eliminated
4319
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    10
           by ordinary simplification. 
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    11
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    12
     and   ! x. (... & x = t & ...) --> P x
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    13
     into  ! x. x = t --> (... & ...) --> P x
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    14
     where the `!x. x=t -->' in the latter formula is eliminated
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    15
           by ordinary simplification.
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    16
11232
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 11221
diff changeset
    17
     And analogously for t=x, but the eqn is not turned around!
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 11221
diff changeset
    18
4319
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    19
     NB Simproc is only triggered by "!x. P(x) & P'(x) --> Q(x)";
38456
6769ccd90ad6 typos in comment
blanchet
parents: 36610
diff changeset
    20
        "!x. x=t --> P(x)" is covered by the congruence rule for -->;
4319
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    21
        "!x. t=x --> P(x)" must be taken care of by an ordinary rewrite rule.
11232
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 11221
diff changeset
    22
        As must be "? x. t=x & P(x)".
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 11221
diff changeset
    23
        
11221
60c6e91f6079 added simproc for bounded quantifiers
nipkow
parents: 7951
diff changeset
    24
     And similarly for the bounded quantifiers.
60c6e91f6079 added simproc for bounded quantifiers
nipkow
parents: 7951
diff changeset
    25
4319
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    26
Gries etc call this the "1 point rules"
31166
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
    27
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
    28
The above also works for !x1..xn. and ?x1..xn by moving the defined
38456
6769ccd90ad6 typos in comment
blanchet
parents: 36610
diff changeset
    29
quantifier inside first, but not for nested bounded quantifiers.
31166
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
    30
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
    31
For set comprehensions the basic permutations
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
    32
      ... & x = t & ...  ->  x = t & (... & ...)
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
    33
      ... & t = x & ...  ->  t = x & (... & ...)
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
    34
are also exported.
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
    35
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
    36
To avoid looping, NONE is returned if the term cannot be rearranged,
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
    37
esp if x=t/t=x sits at the front already.
4319
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    38
*)
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    39
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    40
signature QUANTIFIER1_DATA =
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    41
sig
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    42
  (*abstract syntax*)
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    43
  val dest_eq: term -> (term*term*term)option
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    44
  val dest_conj: term -> (term*term*term)option
11232
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 11221
diff changeset
    45
  val dest_imp:  term -> (term*term*term)option
4319
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    46
  val conj: term
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    47
  val imp:  term
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    48
  (*rules*)
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    49
  val iff_reflection: thm (* P <-> Q ==> P == Q *)
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    50
  val iffI:  thm
12523
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
    51
  val iff_trans: thm
4319
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    52
  val conjI: thm
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    53
  val conjE: thm
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    54
  val impI:  thm
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    55
  val mp:    thm
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    56
  val exI:   thm
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    57
  val exE:   thm
11232
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 11221
diff changeset
    58
  val uncurry: thm (* P --> Q --> R ==> P & Q --> R *)
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 11221
diff changeset
    59
  val iff_allI: thm (* !!x. P x <-> Q x ==> (!x. P x) = (!x. Q x) *)
12523
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
    60
  val iff_exI: thm (* !!x. P x <-> Q x ==> (? x. P x) = (? x. Q x) *)
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
    61
  val all_comm: thm (* (!x y. P x y) = (!y x. P x y) *)
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
    62
  val ex_comm: thm (* (? x y. P x y) = (? y x. P x y) *)
4319
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    63
end;
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    64
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    65
signature QUANTIFIER1 =
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    66
sig
11221
60c6e91f6079 added simproc for bounded quantifiers
nipkow
parents: 7951
diff changeset
    67
  val prove_one_point_all_tac: tactic
60c6e91f6079 added simproc for bounded quantifiers
nipkow
parents: 7951
diff changeset
    68
  val prove_one_point_ex_tac: tactic
17002
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15531
diff changeset
    69
  val rearrange_all: theory -> simpset -> term -> thm option
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15531
diff changeset
    70
  val rearrange_ex:  theory -> simpset -> term -> thm option
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15531
diff changeset
    71
  val rearrange_ball: (simpset -> tactic) -> theory -> simpset -> term -> thm option
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15531
diff changeset
    72
  val rearrange_bex:  (simpset -> tactic) -> theory -> simpset -> term -> thm option
31166
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
    73
  val rearrange_Coll: tactic -> theory -> simpset -> term -> thm option
4319
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    74
end;
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    75
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    76
functor Quantifier1Fun(Data: QUANTIFIER1_DATA): QUANTIFIER1 =
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    77
struct
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    78
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    79
open Data;
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    80
11232
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 11221
diff changeset
    81
(* FIXME: only test! *)
12523
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
    82
fun def xs eq =
31166
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
    83
  (case dest_eq eq of
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
    84
     SOME(c,s,t) =>
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
    85
       let val n = length xs
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
    86
       in s = Bound n andalso not(loose_bvar1(t,n)) orelse
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
    87
          t = Bound n andalso not(loose_bvar1(s,n)) end
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
    88
   | NONE => false);
4319
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
    89
31166
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
    90
fun extract_conj fst xs t = case dest_conj t of NONE => NONE
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
    91
    | SOME(conj,P,Q) =>
31197
c1c163ec6c44 fine-tuned elimination of comprehensions involving x=t.
nipkow
parents: 31166
diff changeset
    92
        (if def xs P then (if fst then NONE else SOME(xs,P,Q)) else
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
    93
         if def xs Q then SOME(xs,Q,P) else
31166
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
    94
         (case extract_conj false xs P of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
    95
            SOME(xs,eq,P') => SOME(xs,eq, conj $ P' $ Q)
31166
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
    96
          | NONE => (case extract_conj false xs Q of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
    97
                       SOME(xs,eq,Q') => SOME(xs,eq,conj $ P $ Q')
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
    98
                     | NONE => NONE)));
11232
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 11221
diff changeset
    99
31166
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
   100
fun extract_imp fst xs t = case dest_imp t of NONE => NONE
31197
c1c163ec6c44 fine-tuned elimination of comprehensions involving x=t.
nipkow
parents: 31166
diff changeset
   101
    | SOME(imp,P,Q) => if def xs P then (if fst then NONE else SOME(xs,P,Q))
31166
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
   102
                       else (case extract_conj false xs P of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
   103
                               SOME(xs,eq,P') => SOME(xs, eq, imp $ P' $ Q)
31166
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
   104
                             | NONE => (case extract_imp false xs Q of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
   105
                                          NONE => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
   106
                                        | SOME(xs,eq,Q') =>
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
   107
                                            SOME(xs,eq,imp$P$Q')));
12523
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   108
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   109
fun extract_quant extract q =
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   110
  let fun exqu xs ((qC as Const(qa,_)) $ Abs(x,T,Q)) =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
   111
            if qa = q then exqu ((qC,x,T)::xs) Q else NONE
31166
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
   112
        | exqu xs P = extract (null xs) xs P
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
   113
  in exqu [] end;
4319
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   114
17002
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15531
diff changeset
   115
fun prove_conv tac thy tu =
38452
abc655166d61 now works for schematic terms as well, thanks to Alex for the `how-to'
nipkow
parents: 36610
diff changeset
   116
  let val ctxt = ProofContext.init_global thy;
abc655166d61 now works for schematic terms as well, thanks to Alex for the `how-to'
nipkow
parents: 36610
diff changeset
   117
      val eq_tu = Logic.mk_equals tu;
abc655166d61 now works for schematic terms as well, thanks to Alex for the `how-to'
nipkow
parents: 36610
diff changeset
   118
      val ([fixed_goal], ctxt') = Variable.import_terms true [eq_tu] ctxt;
abc655166d61 now works for schematic terms as well, thanks to Alex for the `how-to'
nipkow
parents: 36610
diff changeset
   119
      val thm = Goal.prove ctxt' [] [] fixed_goal
abc655166d61 now works for schematic terms as well, thanks to Alex for the `how-to'
nipkow
parents: 36610
diff changeset
   120
            (K (rtac iff_reflection 1 THEN tac));
abc655166d61 now works for schematic terms as well, thanks to Alex for the `how-to'
nipkow
parents: 36610
diff changeset
   121
      val [gen_thm] = Variable.export ctxt' ctxt [thm];
abc655166d61 now works for schematic terms as well, thanks to Alex for the `how-to'
nipkow
parents: 36610
diff changeset
   122
  in gen_thm end;
4319
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   123
12523
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   124
fun qcomm_tac qcomm qI i = REPEAT_DETERM (rtac qcomm i THEN rtac qI i) 
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   125
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   126
(* Proves (? x0..xn. ... & x0 = t & ...) = (? x1..xn x0. x0 = t & ... & ...)
11221
60c6e91f6079 added simproc for bounded quantifiers
nipkow
parents: 7951
diff changeset
   127
   Better: instantiate exI
60c6e91f6079 added simproc for bounded quantifiers
nipkow
parents: 7951
diff changeset
   128
*)
12523
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   129
local
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   130
val excomm = ex_comm RS iff_trans
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   131
in
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   132
val prove_one_point_ex_tac = qcomm_tac excomm iff_exI 1 THEN rtac iffI 1 THEN
11221
60c6e91f6079 added simproc for bounded quantifiers
nipkow
parents: 7951
diff changeset
   133
    ALLGOALS(EVERY'[etac exE, REPEAT_DETERM o (etac conjE), rtac exI,
12523
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   134
                    DEPTH_SOLVE_1 o (ares_tac [conjI])])
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   135
end;
11221
60c6e91f6079 added simproc for bounded quantifiers
nipkow
parents: 7951
diff changeset
   136
12523
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   137
(* Proves (! x0..xn. (... & x0 = t & ...) --> P x0) =
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   138
          (! x1..xn x0. x0 = t --> (... & ...) --> P x0)
11221
60c6e91f6079 added simproc for bounded quantifiers
nipkow
parents: 7951
diff changeset
   139
*)
11232
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 11221
diff changeset
   140
local
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 11221
diff changeset
   141
val tac = SELECT_GOAL
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 11221
diff changeset
   142
          (EVERY1[REPEAT o (dtac uncurry), REPEAT o (rtac impI), etac mp,
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 11221
diff changeset
   143
                  REPEAT o (etac conjE), REPEAT o (ares_tac [conjI])])
12523
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   144
val allcomm = all_comm RS iff_trans
11232
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 11221
diff changeset
   145
in
12523
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   146
val prove_one_point_all_tac =
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   147
      EVERY1[qcomm_tac allcomm iff_allI,rtac iff_allI, rtac iffI, tac, tac]
11232
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 11221
diff changeset
   148
end
4319
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   149
12523
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   150
fun renumber l u (Bound i) = Bound(if i < l orelse i > u then i else
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   151
                                   if i=u then l else i+1)
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   152
  | renumber l u (s$t) = renumber l u s $ renumber l u t
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   153
  | renumber l u (Abs(x,T,t)) = Abs(x,T,renumber (l+1) (u+1) t)
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   154
  | renumber _ _ atom = atom;
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   155
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   156
fun quantify qC x T xs P =
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   157
  let fun quant [] P = P
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   158
        | quant ((qC,x,T)::xs) P = quant xs (qC $ Abs(x,T,P))
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   159
      val n = length xs
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   160
      val Q = if n=0 then P else renumber 0 n P
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   161
  in quant xs (qC $ Abs(x,T,Q)) end;
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   162
17002
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15531
diff changeset
   163
fun rearrange_all thy _ (F as (all as Const(q,_)) $ Abs(x,T, P)) =
31166
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
   164
     (case extract_quant extract_imp q P of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
   165
        NONE => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
   166
      | SOME(xs,eq,Q) =>
12523
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   167
          let val R = quantify all x T xs (imp $ eq $ Q)
17002
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15531
diff changeset
   168
          in SOME(prove_conv prove_one_point_all_tac thy (F,R)) end)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
   169
  | rearrange_all _ _ _ = NONE;
4319
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   170
17002
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15531
diff changeset
   171
fun rearrange_ball tac thy ss (F as Ball $ A $ Abs(x,T,P)) =
31166
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
   172
     (case extract_imp true [] P of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
   173
        NONE => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
   174
      | SOME(xs,eq,Q) => if not(null xs) then NONE else
11232
558a4feebb04 generalization of 1 point rules for ALL
nipkow
parents: 11221
diff changeset
   175
          let val R = imp $ eq $ Q
17002
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15531
diff changeset
   176
          in SOME(prove_conv (tac ss) thy (F,Ball $ A $ Abs(x,T,R))) end)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
   177
  | rearrange_ball _ _ _ _ = NONE;
4319
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   178
17002
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15531
diff changeset
   179
fun rearrange_ex thy _ (F as (ex as Const(q,_)) $ Abs(x,T,P)) =
31166
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
   180
     (case extract_quant extract_conj q P of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
   181
        NONE => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
   182
      | SOME(xs,eq,Q) =>
12523
0d8d5bf549b0 now permutations of quantifiers are allowed as well.
nipkow
parents: 11232
diff changeset
   183
          let val R = quantify ex x T xs (conj $ eq $ Q)
17002
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15531
diff changeset
   184
          in SOME(prove_conv prove_one_point_ex_tac thy (F,R)) end)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
   185
  | rearrange_ex _ _ _ = NONE;
4319
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   186
17002
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15531
diff changeset
   187
fun rearrange_bex tac thy ss (F as Bex $ A $ Abs(x,T,P)) =
31166
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
   188
     (case extract_conj true [] P of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
   189
        NONE => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
   190
      | SOME(xs,eq,Q) => if not(null xs) then NONE else
17002
fb9261990ffe simprocs: Simplifier.inherit_bounds;
wenzelm
parents: 15531
diff changeset
   191
          SOME(prove_conv (tac ss) thy (F,Bex $ A $ Abs(x,T,conj$eq$Q))))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
   192
  | rearrange_bex _ _ _ _ = NONE;
11221
60c6e91f6079 added simproc for bounded quantifiers
nipkow
parents: 7951
diff changeset
   193
31166
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
   194
fun rearrange_Coll tac thy _ (F as Coll $ Abs(x,T,P)) =
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
   195
     (case extract_conj true [] P of
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
   196
        NONE => NONE
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
   197
      | SOME(_,eq,Q) =>
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
   198
          let val R = Coll $ Abs(x,T, conj $ eq $ Q)
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
   199
          in SOME(prove_conv tac thy (F,R)) end);
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 20049
diff changeset
   200
4319
afb60b8bf15e Quantifier elimination procs.
nipkow
parents:
diff changeset
   201
end;