src/HOL/Algebra/abstract/Factor.thy
author chaieb
Wed, 19 May 2004 11:23:59 +0200
changeset 14758 af3b71a46a1c
parent 13735 7de9342aca7a
child 17479 68a7acb5f22e
permissions -rw-r--r--
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller. the tactic has also changed and allows the abstaction over fuction occurences whose type is nat or int.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     1
(*
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     2
    Factorisation within a factorial domain
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     3
    $Id$
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     4
    Author: Clemens Ballarin, started 25 November 1997
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     5
*)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     6
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     7
Factor = Ring +
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     8
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
     9
consts
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 7998
diff changeset
    10
  Factorisation	:: ['a::ring, 'a list, 'a] => bool
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    11
  (* factorisation of x into a list of irred factors and a unit u *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    12
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    13
defs
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    14
  Factorisation_def	"Factorisation x factors u == 
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    15
                           x = foldr op* factors u &
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 7998
diff changeset
    16
                           (ALL a : set factors. irred a) & u dvd 1"
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    17
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    18
end