src/HOL/Decision_Procs/commutative_ring_tac.ML
author wenzelm
Fri, 06 Mar 2015 23:33:25 +0100
changeset 59629 0d77c51b5040
parent 59621 291934bac95e
child 60143 2cd31c81e0e7
permissions -rw-r--r--
clarified context;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37744
3daaf23b9ab4 tuned titles
haftmann
parents: 35267
diff changeset
     1
(*  Title:      HOL/Decision_Procs/commutative_ring_tac.ML
3daaf23b9ab4 tuned titles
haftmann
parents: 35267
diff changeset
     2
    Author:     Amine Chaieb
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     3
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     4
Tactic for solving equalities over commutative rings.
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     5
*)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     6
33356
9157d0f9f00e moved Commutative_Ring into session Decision_Procs
haftmann
parents: 32149
diff changeset
     7
signature COMMUTATIVE_RING_TAC =
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     8
sig
33356
9157d0f9f00e moved Commutative_Ring into session Decision_Procs
haftmann
parents: 32149
diff changeset
     9
  val tac: Proof.context -> int -> tactic
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    10
end
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    11
33356
9157d0f9f00e moved Commutative_Ring into session Decision_Procs
haftmann
parents: 32149
diff changeset
    12
structure Commutative_Ring_Tac: COMMUTATIVE_RING_TAC =
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    13
struct
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    14
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    15
(* Zero and One of the commutative ring *)
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 34974
diff changeset
    16
fun cring_zero T = Const (@{const_name Groups.zero}, T);
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 34974
diff changeset
    17
fun cring_one T = Const (@{const_name Groups.one}, T);
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    18
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    19
(* reification functions *)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    20
(* add two polynom expressions *)
33356
9157d0f9f00e moved Commutative_Ring into session Decision_Procs
haftmann
parents: 32149
diff changeset
    21
fun polT t = Type (@{type_name Commutative_Ring.pol}, [t]);
9157d0f9f00e moved Commutative_Ring into session Decision_Procs
haftmann
parents: 32149
diff changeset
    22
fun polexT t = Type (@{type_name Commutative_Ring.polex}, [t]);
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    23
20623
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
    24
(* pol *)
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    25
fun pol_Pc t =
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    26
  Const (@{const_name Commutative_Ring.pol.Pc}, t --> polT t);
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    27
fun pol_Pinj t =
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    28
  Const (@{const_name Commutative_Ring.pol.Pinj}, HOLogic.natT --> polT t --> polT t);
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    29
fun pol_PX t =
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    30
  Const (@{const_name Commutative_Ring.pol.PX}, polT t --> HOLogic.natT --> polT t --> polT t);
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    31
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    32
(* polex *)
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    33
fun polex_add t =
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    34
  Const (@{const_name Commutative_Ring.polex.Add}, polexT t --> polexT t --> polexT t);
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    35
fun polex_sub t =
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    36
  Const (@{const_name Commutative_Ring.polex.Sub}, polexT t --> polexT t --> polexT t);
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    37
fun polex_mul t =
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    38
  Const (@{const_name Commutative_Ring.polex.Mul}, polexT t --> polexT t --> polexT t);
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    39
fun polex_neg t =
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    40
  Const (@{const_name Commutative_Ring.polex.Neg}, polexT t --> polexT t);
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    41
fun polex_pol t =
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    42
  Const (@{const_name Commutative_Ring.polex.Pol}, polT t --> polexT t);
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    43
fun polex_pow t =
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    44
  Const (@{const_name Commutative_Ring.polex.Pow}, polexT t --> HOLogic.natT --> polexT t);
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    45
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    46
(* reification of polynoms : primitive cring expressions *)
22950
haftmann
parents: 21588
diff changeset
    47
fun reif_pol T vs (t as Free _) =
haftmann
parents: 21588
diff changeset
    48
      let
haftmann
parents: 21588
diff changeset
    49
        val one = @{term "1::nat"};
31986
a68f88d264f7 dropped find_index_eq
haftmann
parents: 31242
diff changeset
    50
        val i = find_index (fn t' => t' = t) vs
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    51
      in
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    52
        if i = 0 then
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    53
          pol_PX T $ (pol_Pc T $ cring_one T) $ one $ (pol_Pc T $ cring_zero T)
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    54
        else
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    55
          pol_Pinj T $ HOLogic.mk_nat i $
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    56
            (pol_PX T $ (pol_Pc T $ cring_one T) $ one $ (pol_Pc T $ cring_zero T))
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    57
        end
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    58
  | reif_pol T _ t = pol_Pc T $ t;
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    59
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    60
(* reification of polynom expressions *)
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 34974
diff changeset
    61
fun reif_polex T vs (Const (@{const_name Groups.plus}, _) $ a $ b) =
22950
haftmann
parents: 21588
diff changeset
    62
      polex_add T $ reif_polex T vs a $ reif_polex T vs b
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 34974
diff changeset
    63
  | reif_polex T vs (Const (@{const_name Groups.minus}, _) $ a $ b) =
22950
haftmann
parents: 21588
diff changeset
    64
      polex_sub T $ reif_polex T vs a $ reif_polex T vs b
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 34974
diff changeset
    65
  | reif_polex T vs (Const (@{const_name Groups.times}, _) $ a $ b) =
22950
haftmann
parents: 21588
diff changeset
    66
      polex_mul T $ reif_polex T vs a $ reif_polex T vs b
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 34974
diff changeset
    67
  | reif_polex T vs (Const (@{const_name Groups.uminus}, _) $ a) =
22950
haftmann
parents: 21588
diff changeset
    68
      polex_neg T $ reif_polex T vs a
24996
ebd5f4cc7118 moved class power to theory Power
haftmann
parents: 24630
diff changeset
    69
  | reif_polex T vs (Const (@{const_name Power.power}, _) $ a $ n) =
22950
haftmann
parents: 21588
diff changeset
    70
      polex_pow T $ reif_polex T vs a $ n
haftmann
parents: 21588
diff changeset
    71
  | reif_polex T vs t = polex_pol T $ reif_pol T vs t;
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    72
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    73
(* reification of the equation *)
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    74
val cr_sort = @{sort comm_ring_1};
22950
haftmann
parents: 21588
diff changeset
    75
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    76
fun reif_eq ctxt (eq as Const (@{const_name HOL.eq}, Type (@{type_name fun}, [T, _])) $ lhs $ rhs) =
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    77
      if Sign.of_sort (Proof_Context.theory_of ctxt) (T, cr_sort) then
22950
haftmann
parents: 21588
diff changeset
    78
        let
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
    79
          val fs = Misc_Legacy.term_frees eq;
59629
0d77c51b5040 clarified context;
wenzelm
parents: 59621
diff changeset
    80
          val cvs = Thm.cterm_of ctxt (HOLogic.mk_list T fs);
0d77c51b5040 clarified context;
wenzelm
parents: 59621
diff changeset
    81
          val clhs = Thm.cterm_of ctxt (reif_polex T fs lhs);
0d77c51b5040 clarified context;
wenzelm
parents: 59621
diff changeset
    82
          val crhs = Thm.cterm_of ctxt (reif_polex T fs rhs);
0d77c51b5040 clarified context;
wenzelm
parents: 59621
diff changeset
    83
          val ca = Thm.ctyp_of ctxt T;
22950
haftmann
parents: 21588
diff changeset
    84
        in (ca, cvs, clhs, crhs) end
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    85
      else error ("reif_eq: not an equation over " ^ Syntax.string_of_sort ctxt cr_sort)
33356
9157d0f9f00e moved Commutative_Ring into session Decision_Procs
haftmann
parents: 32149
diff changeset
    86
  | reif_eq _ _ = error "reif_eq: not an equation";
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    87
22950
haftmann
parents: 21588
diff changeset
    88
(* The cring tactic *)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    89
(* Attention: You have to make sure that no t^0 is in the goal!! *)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    90
(* Use simply rewriting t^0 = 1 *)
20623
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
    91
val cring_simps =
58645
94bef115c08f more foundational definition for predicate even
haftmann
parents: 55793
diff changeset
    92
  @{thms mkPX_def mkPinj_def sub_def power_add even_iff_mod_2_eq_zero pow_if power_add [symmetric]};
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    93
33356
9157d0f9f00e moved Commutative_Ring into session Decision_Procs
haftmann
parents: 32149
diff changeset
    94
fun tac ctxt = SUBGOAL (fn (g, i) =>
20623
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
    95
  let
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
    96
    val cring_ctxt = ctxt addsimps cring_simps;  (*FIXME really the full simpset!?*)
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    97
    val (ca, cvs, clhs, crhs) = reif_eq ctxt (HOLogic.dest_Trueprop g);
20623
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
    98
    val norm_eq_th =
55793
52c8f934ea6f tuned whitespace;
wenzelm
parents: 51717
diff changeset
    99
      simplify cring_ctxt (instantiate' [SOME ca] [SOME clhs, SOME crhs, SOME cvs] @{thm norm_eq});
20623
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
   100
  in
46708
b138dee7bed3 prefer cut_tac, where it is clear that the special variants cut_rules_tac or cut_facts_tac are not required;
wenzelm
parents: 44121
diff changeset
   101
    cut_tac norm_eq_th i
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
   102
    THEN (simp_tac cring_ctxt i)
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47432
diff changeset
   103
    THEN (simp_tac cring_ctxt i)
20623
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
   104
  end);
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
   105
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   106
end;