src/HOL/Library/comm_ring.ML
author haftmann
Thu, 12 Mar 2009 18:01:25 +0100
changeset 30495 a5f1e4f46d14
parent 29265 5b4247055bd7
child 30510 4120fc59dd85
permissions -rw-r--r--
tuned
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
29265
5b4247055bd7 moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents: 26939
diff changeset
     1
(*  Author:     Amine Chaieb
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     2
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     3
Tactic for solving equalities over commutative rings.
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     4
*)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     5
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     6
signature COMM_RING =
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     7
sig
20623
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
     8
  val comm_ring_tac : Proof.context -> int -> tactic
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 17516
diff changeset
     9
  val setup : theory -> theory
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
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    12
structure CommRing: COMM_RING =
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
(* The Cring exception for erronous uses of cring_tac *)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    16
exception CRing of string;
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    17
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    18
(* Zero and One of the commutative ring *)
22997
d4f3b015b50b canonical prefixing of class constants
haftmann
parents: 22963
diff changeset
    19
fun cring_zero T = Const (@{const_name HOL.zero}, T);
d4f3b015b50b canonical prefixing of class constants
haftmann
parents: 22963
diff changeset
    20
fun cring_one T = Const (@{const_name HOL.one}, T);
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    21
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    22
(* reification functions *)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    23
(* add two polynom expressions *)
22950
haftmann
parents: 21588
diff changeset
    24
fun polT t = Type ("Commutative_Ring.pol", [t]);
haftmann
parents: 21588
diff changeset
    25
fun polexT t = Type ("Commutative_Ring.polex", [t]);
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    26
20623
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
    27
(* pol *)
22950
haftmann
parents: 21588
diff changeset
    28
fun pol_Pc t = Const ("Commutative_Ring.pol.Pc", t --> polT t);
haftmann
parents: 21588
diff changeset
    29
fun pol_Pinj t = Const ("Commutative_Ring.pol.Pinj", HOLogic.natT --> polT t --> polT t);
haftmann
parents: 21588
diff changeset
    30
fun pol_PX t = Const ("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 *)
22950
haftmann
parents: 21588
diff changeset
    33
fun polex_add t = Const ("Commutative_Ring.polex.Add", polexT t --> polexT t --> polexT t);
haftmann
parents: 21588
diff changeset
    34
fun polex_sub t = Const ("Commutative_Ring.polex.Sub", polexT t --> polexT t --> polexT t);
haftmann
parents: 21588
diff changeset
    35
fun polex_mul t = Const ("Commutative_Ring.polex.Mul", polexT t --> polexT t --> polexT t);
haftmann
parents: 21588
diff changeset
    36
fun polex_neg t = Const ("Commutative_Ring.polex.Neg", polexT t --> polexT t);
haftmann
parents: 21588
diff changeset
    37
fun polex_pol t = Const ("Commutative_Ring.polex.Pol", polT t --> polexT t);
haftmann
parents: 21588
diff changeset
    38
fun polex_pow t = Const ("Commutative_Ring.polex.Pow", polexT t --> HOLogic.natT --> polexT t);
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    39
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    40
(* reification of polynoms : primitive cring expressions *)
22950
haftmann
parents: 21588
diff changeset
    41
fun reif_pol T vs (t as Free _) =
haftmann
parents: 21588
diff changeset
    42
      let
haftmann
parents: 21588
diff changeset
    43
        val one = @{term "1::nat"};
haftmann
parents: 21588
diff changeset
    44
        val i = find_index_eq t vs
haftmann
parents: 21588
diff changeset
    45
      in if i = 0
haftmann
parents: 21588
diff changeset
    46
        then pol_PX T $ (pol_Pc T $ cring_one T)
haftmann
parents: 21588
diff changeset
    47
          $ one $ (pol_Pc T $ cring_zero T)
24630
351a308ab58d simplified type int (eliminated IntInf.int, integer);
wenzelm
parents: 23261
diff changeset
    48
        else pol_Pinj T $ HOLogic.mk_nat i
22950
haftmann
parents: 21588
diff changeset
    49
          $ (pol_PX T $ (pol_Pc T $ cring_one T)
haftmann
parents: 21588
diff changeset
    50
            $ one $ (pol_Pc T $ cring_zero T))
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    51
        end
22950
haftmann
parents: 21588
diff changeset
    52
  | reif_pol T vs t = pol_Pc T $ t;
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    53
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    54
(* reification of polynom expressions *)
22997
d4f3b015b50b canonical prefixing of class constants
haftmann
parents: 22963
diff changeset
    55
fun reif_polex T vs (Const (@{const_name HOL.plus}, _) $ a $ b) =
22950
haftmann
parents: 21588
diff changeset
    56
      polex_add T $ reif_polex T vs a $ reif_polex T vs b
22997
d4f3b015b50b canonical prefixing of class constants
haftmann
parents: 22963
diff changeset
    57
  | reif_polex T vs (Const (@{const_name HOL.minus}, _) $ a $ b) =
22950
haftmann
parents: 21588
diff changeset
    58
      polex_sub T $ reif_polex T vs a $ reif_polex T vs b
22997
d4f3b015b50b canonical prefixing of class constants
haftmann
parents: 22963
diff changeset
    59
  | reif_polex T vs (Const (@{const_name HOL.times}, _) $ a $ b) =
22950
haftmann
parents: 21588
diff changeset
    60
      polex_mul T $ reif_polex T vs a $ reif_polex T vs b
22997
d4f3b015b50b canonical prefixing of class constants
haftmann
parents: 22963
diff changeset
    61
  | reif_polex T vs (Const (@{const_name HOL.uminus}, _) $ a) =
22950
haftmann
parents: 21588
diff changeset
    62
      polex_neg T $ reif_polex T vs a
24996
ebd5f4cc7118 moved class power to theory Power
haftmann
parents: 24630
diff changeset
    63
  | reif_polex T vs (Const (@{const_name Power.power}, _) $ a $ n) =
22950
haftmann
parents: 21588
diff changeset
    64
      polex_pow T $ reif_polex T vs a $ n
haftmann
parents: 21588
diff changeset
    65
  | 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
    66
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    67
(* reification of the equation *)
22950
haftmann
parents: 21588
diff changeset
    68
val TFree (_, cr_sort) = @{typ "'a :: {comm_ring, recpower}"};
haftmann
parents: 21588
diff changeset
    69
haftmann
parents: 21588
diff changeset
    70
fun reif_eq thy (eq as Const("op =", Type("fun", [T, _])) $ lhs $ rhs) =
haftmann
parents: 21588
diff changeset
    71
      if Sign.of_sort thy (T, cr_sort) then
haftmann
parents: 21588
diff changeset
    72
        let
29265
5b4247055bd7 moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents: 26939
diff changeset
    73
          val fs = OldTerm.term_frees eq;
22950
haftmann
parents: 21588
diff changeset
    74
          val cvs = cterm_of thy (HOLogic.mk_list T fs);
haftmann
parents: 21588
diff changeset
    75
          val clhs = cterm_of thy (reif_polex T fs lhs);
haftmann
parents: 21588
diff changeset
    76
          val crhs = cterm_of thy (reif_polex T fs rhs);
haftmann
parents: 21588
diff changeset
    77
          val ca = ctyp_of thy T;
haftmann
parents: 21588
diff changeset
    78
        in (ca, cvs, clhs, crhs) end
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 24996
diff changeset
    79
      else raise CRing ("reif_eq: not an equation over " ^ Syntax.string_of_sort_global thy cr_sort)
20623
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
    80
  | reif_eq _ _ = raise CRing "reif_eq: not an equation";
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    81
22950
haftmann
parents: 21588
diff changeset
    82
(* The cring tactic *)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    83
(* 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
    84
(* Use simply rewriting t^0 = 1 *)
20623
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
    85
val cring_simps =
22950
haftmann
parents: 21588
diff changeset
    86
  [@{thm mkPX_def}, @{thm mkPinj_def}, @{thm sub_def}, @{thm power_add},
haftmann
parents: 21588
diff changeset
    87
    @{thm even_def}, @{thm pow_if}, sym OF [@{thm power_add}]];
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    88
20623
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
    89
fun comm_ring_tac ctxt = SUBGOAL (fn (g, i) =>
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
    90
  let
22950
haftmann
parents: 21588
diff changeset
    91
    val thy = ProofContext.theory_of ctxt;
haftmann
parents: 21588
diff changeset
    92
    val cring_ss = Simplifier.local_simpset_of ctxt  (*FIXME really the full simpset!?*)
haftmann
parents: 21588
diff changeset
    93
      addsimps cring_simps;
20623
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
    94
    val (ca, cvs, clhs, crhs) = reif_eq thy (HOLogic.dest_Trueprop g)
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
    95
    val norm_eq_th =
22950
haftmann
parents: 21588
diff changeset
    96
      simplify cring_ss (instantiate' [SOME ca] [SOME clhs, SOME crhs, SOME cvs] @{thm norm_eq})
20623
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
    97
  in
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
    98
    cut_rules_tac [norm_eq_th] i
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
    99
    THEN (simp_tac cring_ss i)
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
   100
    THEN (simp_tac cring_ss i)
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
   101
  end);
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
   102
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
   103
val comm_ring_meth =
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 20713
diff changeset
   104
  Method.ctxt_args (Method.SIMPLE_METHOD' o comm_ring_tac);
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   105
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   106
val setup =
20623
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
   107
  Method.add_method ("comm_ring", comm_ring_meth,
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
   108
    "reflective decision procedure for equalities over commutative rings") #>
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
   109
  Method.add_method ("algebra", comm_ring_meth,
6ae83d153dd4 tuned method setup;
wenzelm
parents: 19233
diff changeset
   110
    "method for proving algebraic properties (same as comm_ring)");
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   111
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   112
end;