src/HOL/Algebra/abstract/Ideal.ML
author kleing
Tue, 13 May 2003 08:59:21 +0200
changeset 14024 213dcc39358f
parent 13735 7de9342aca7a
child 17479 68a7acb5f22e
permissions -rw-r--r--
HOL-Real -> HOL-Complex
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
    Ideals for commutative rings
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 24 September 1999
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
(* is_ideal *)
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
Goalw [is_ideal_def]
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    10
  "!! I. [| !! a b. [| a:I; b:I |] ==> a + b : I; \
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10833
diff changeset
    11
\     !! a. a:I ==> - a : I; 0 : I; \
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    12
\     !! a r. a:I ==> a * r : I |] ==> is_ideal I";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    13
by Auto_tac;
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    14
qed "is_idealI";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    15
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    16
Goalw [is_ideal_def] "[| is_ideal I; a:I; b:I |] ==> a + b : I";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    17
by (Fast_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    18
qed "is_ideal_add";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    19
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    20
Goalw [is_ideal_def] "[| is_ideal I; a:I |] ==> - a : I";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    21
by (Fast_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    22
qed "is_ideal_uminus";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    23
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10833
diff changeset
    24
Goalw [is_ideal_def] "[| is_ideal I |] ==> 0 : I";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    25
by (Fast_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    26
qed "is_ideal_zero";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    27
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    28
Goalw [is_ideal_def] "[| is_ideal I; a:I |] ==> a * r : I";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    29
by (Fast_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    30
qed "is_ideal_mult";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    31
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    32
Goalw [dvd_def, is_ideal_def] "[| a dvd b; is_ideal I; a:I |] ==> b:I";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    33
by (Fast_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    34
qed "is_ideal_dvd";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    35
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    36
Goalw [is_ideal_def] "is_ideal (UNIV::('a::ring) set)";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    37
by Auto_tac;
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    38
qed "UNIV_is_ideal";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    39
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10833
diff changeset
    40
Goalw [is_ideal_def] "is_ideal {0::'a::ring}";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    41
by Auto_tac;
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    42
qed "zero_is_ideal";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    43
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    44
Addsimps [is_ideal_add, is_ideal_uminus, is_ideal_zero, is_ideal_mult,
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    45
  UNIV_is_ideal, zero_is_ideal];
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    46
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    47
Goal "is_ideal {x::'a::ring. a dvd x}";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    48
by (rtac is_idealI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    49
by Auto_tac;
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    50
qed "is_ideal_1";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    51
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    52
Goal "is_ideal {x::'a::ring. EX u v. x = a * u + b * v}";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    53
by (rtac is_idealI 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
    54
(* by Auto_tac; FIXME: makes Zumkeller's order fail (raises exn Domain) *)
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
    55
by (Clarify_tac 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
    56
by (Clarify_tac 2);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
    57
by (Clarify_tac 3);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
    58
by (Clarify_tac 4);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    59
by (res_inst_tac [("x", "u+ua")] exI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    60
by (res_inst_tac [("x", "v+va")] exI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    61
by (res_inst_tac [("x", "-u")] exI 2);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    62
by (res_inst_tac [("x", "-v")] exI 2);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10833
diff changeset
    63
by (res_inst_tac [("x", "0")] exI 3);
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10833
diff changeset
    64
by (res_inst_tac [("x", "0")] exI 3);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    65
by (res_inst_tac [("x", "u * r")] exI 4);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    66
by (res_inst_tac [("x", "v * r")] exI 4);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
    67
by (REPEAT (Simp_tac 1));
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    68
qed "is_ideal_2";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    69
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    70
(* ideal_of *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    71
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    72
Goalw [is_ideal_def, ideal_of_def] "is_ideal (ideal_of S)";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    73
by (Blast_tac 1);  (* Here, blast_tac is much superior to fast_tac! *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    74
qed "ideal_of_is_ideal";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    75
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    76
Goalw [ideal_of_def] "a:S ==> a : (ideal_of S)";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    77
by Auto_tac;
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    78
qed "generator_in_ideal";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    79
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
    80
Goalw [ideal_of_def] "ideal_of {1::'a::ring} = UNIV";
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
    81
by (force_tac (claset() addDs [is_ideal_mult], 
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
    82
  simpset() addsimps [l_one] delsimprocs [ring_simproc]) 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
    83
  (* FIXME: Zumkeller's order raises Domain exn *)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    84
qed "ideal_of_one_eq";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    85
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10833
diff changeset
    86
Goalw [ideal_of_def] "ideal_of {} = {0::'a::ring}";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    87
by (rtac subset_antisym 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    88
by (rtac subsetI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    89
by (dtac InterD 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    90
by (assume_tac 2);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    91
by (auto_tac (claset(), simpset() addsimps [is_ideal_zero]));
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    92
qed "ideal_of_empty_eq";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    93
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    94
Goalw [ideal_of_def] "ideal_of {a} = {x::'a::ring. a dvd x}";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    95
by (rtac subset_antisym 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    96
by (rtac subsetI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    97
by (dtac InterD 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    98
by (assume_tac 2);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
    99
by (auto_tac (claset() addIs [is_ideal_1], simpset()));
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   100
by (asm_simp_tac (simpset() addsimps [is_ideal_dvd]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   101
qed "pideal_structure";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   102
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   103
Goalw [ideal_of_def]
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   104
  "ideal_of {a, b} = {x::'a::ring. EX u v. x = a * u + b * v}";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   105
by (rtac subset_antisym 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   106
by (rtac subsetI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   107
by (dtac InterD 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   108
by (assume_tac 2);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
   109
by (auto_tac (claset() addIs [is_ideal_2],
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
   110
  simpset() delsimprocs [ring_simproc]));
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
   111
(* FIXME: Zumkeller's order *)
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
   112
by (res_inst_tac [("x", "1")] exI 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10833
diff changeset
   113
by (res_inst_tac [("x", "0")] exI 1);
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10833
diff changeset
   114
by (res_inst_tac [("x", "0")] exI 2);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
   115
by (res_inst_tac [("x", "1")] exI 2);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   116
by (Simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   117
by (Simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   118
qed "ideal_of_2_structure";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   119
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   120
Goalw [ideal_of_def] "A <= B ==> ideal_of A <= ideal_of B";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   121
by Auto_tac;
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   122
qed "ideal_of_mono";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   123
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10833
diff changeset
   124
Goal "ideal_of {0::'a::ring} = {0}";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   125
by (simp_tac (simpset() addsimps [pideal_structure]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   126
by (rtac subset_antisym 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   127
by (auto_tac (claset() addIs [dvd_zero_left], simpset()));
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   128
qed "ideal_of_zero_eq";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   129
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   130
Goal "[| is_ideal I; a : I |] ==> ideal_of {a::'a::ring} <= I";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   131
by (auto_tac (claset(),
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   132
  simpset() addsimps [pideal_structure, is_ideal_dvd]));
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   133
qed "element_generates_subideal";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   134
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   135
(* is_pideal *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   136
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   137
Goalw [is_pideal_def] "is_pideal (I::('a::ring) set) ==> is_ideal I";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   138
by (fast_tac (claset() addIs [ideal_of_is_ideal]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   139
qed "is_pideal_imp_is_ideal";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   140
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   141
Goalw [is_pideal_def] "is_pideal (ideal_of {a::'a::ring})";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   142
by (Fast_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   143
qed "pideal_is_pideal";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   144
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   145
Goalw [is_pideal_def] "is_pideal I ==> EX a. I = ideal_of {a}";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   146
by (assume_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   147
qed "is_pidealD";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   148
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   149
(* Ideals and divisibility *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   150
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   151
Goal "b dvd a ==> ideal_of {a::'a::ring} <= ideal_of {b}";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   152
by (auto_tac (claset() addIs [dvd_trans_ring],
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   153
  simpset() addsimps [pideal_structure]));
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   154
qed "dvd_imp_subideal";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   155
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   156
Goal "ideal_of {a::'a::ring} <= ideal_of {b} ==> b dvd a";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   157
by (auto_tac (claset() addSDs [subsetD],
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   158
  simpset() addsimps [pideal_structure]));
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   159
qed "subideal_imp_dvd";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   160
10789
260fa2c67e3e Changed priority of dvd from 70 to 50 as befits a relation.
nipkow
parents: 10230
diff changeset
   161
Goal "(ideal_of {a::'a::ring} <= ideal_of {b}) = (b dvd a)";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   162
by (rtac iffI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   163
by (REPEAT (ares_tac [subideal_imp_dvd, dvd_imp_subideal] 1));
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   164
qed "subideal_is_dvd";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   165
10789
260fa2c67e3e Changed priority of dvd from 70 to 50 as befits a relation.
nipkow
parents: 10230
diff changeset
   166
Goal "(ideal_of {a::'a::ring} < ideal_of {b}) ==> ~ a dvd b";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   167
by (full_simp_tac (simpset() addsimps [psubset_eq, pideal_structure]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   168
by (etac conjE 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   169
by (dres_inst_tac [("c", "a")] subsetD 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   170
by (auto_tac (claset() addIs [dvd_trans_ring],
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   171
  simpset()));
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   172
qed "psubideal_not_dvd";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   173
10789
260fa2c67e3e Changed priority of dvd from 70 to 50 as befits a relation.
nipkow
parents: 10230
diff changeset
   174
Goal "[| b dvd a; ~ a dvd b |] ==> ideal_of {a::'a::ring} < ideal_of {b}";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   175
by (rtac psubsetI 1);
10198
paulson
parents: 8707
diff changeset
   176
by (etac dvd_imp_subideal 1);
paulson
parents: 8707
diff changeset
   177
by (blast_tac (claset() addIs [dvd_imp_subideal, subideal_imp_dvd]) 1); 
11384
cde6edd51ff6 resolved name clash
paulson
parents: 11093
diff changeset
   178
qed "not_dvd_psubideal_singleton";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   179
10789
260fa2c67e3e Changed priority of dvd from 70 to 50 as befits a relation.
nipkow
parents: 10230
diff changeset
   180
Goal "(ideal_of {a::'a::ring} < ideal_of {b}) = (b dvd a & ~ a dvd b)";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   181
by (rtac iffI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   182
by (REPEAT (ares_tac
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   183
  [conjI, psubideal_not_dvd, psubset_imp_subset RS subideal_imp_dvd] 1));
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   184
by (etac conjE 1);
11384
cde6edd51ff6 resolved name clash
paulson
parents: 11093
diff changeset
   185
by (REPEAT (ares_tac [not_dvd_psubideal_singleton] 1));
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   186
qed "psubideal_is_dvd";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   187
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   188
Goal "[| a dvd b; b dvd a |] ==> ideal_of {a::'a::ring} = ideal_of {b}";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   189
by (rtac subset_antisym 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   190
by (REPEAT (ares_tac [dvd_imp_subideal] 1));
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   191
qed "assoc_pideal_eq";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   192
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   193
AddIffs [subideal_is_dvd, psubideal_is_dvd];
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   194
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   195
Goal "!!a::'a::ring. a dvd b ==> b : (ideal_of {a})";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   196
by (rtac is_ideal_dvd 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   197
by (assume_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   198
by (rtac ideal_of_is_ideal 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   199
by (rtac generator_in_ideal 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   200
by (Simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   201
qed "dvd_imp_in_pideal";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   202
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   203
Goal "!!a::'a::ring. b : (ideal_of {a}) ==> a dvd b";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   204
by (full_simp_tac (simpset() addsimps [pideal_structure]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   205
qed "in_pideal_imp_dvd";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   206
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   207
Goal "~ (a dvd b) ==> ideal_of {a::'a::ring} < ideal_of {a, b}";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   208
by (asm_simp_tac (simpset() addsimps [psubset_eq, ideal_of_mono]) 1);
10230
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   209
by (etac contrapos_pp 1);
5eb935d6cc69 tidying and renaming of contrapos rules
paulson
parents: 10198
diff changeset
   210
by (full_simp_tac (simpset() addsimps [ideal_of_2_structure]) 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   211
by (rtac in_pideal_imp_dvd 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   212
by (Asm_simp_tac 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10833
diff changeset
   213
by (res_inst_tac [("x", "0")] exI 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
   214
by (res_inst_tac [("x", "1")] exI 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   215
by (Simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   216
qed "not_dvd_psubideal";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   217
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
   218
Goalw [thm "irred_def"]
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   219
  "[| irred (a::'a::ring); is_pideal I; ideal_of {a} < I |] ==> x : I";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   220
by (dtac is_pidealD 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   221
by (etac exE 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   222
by (Clarify_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   223
by (eres_inst_tac [("x", "aa")] allE 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   224
by (Clarify_tac 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
   225
by (dres_inst_tac [("a", "1")] dvd_imp_subideal 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   226
by (auto_tac (claset(), simpset() addsimps [ideal_of_one_eq]));
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   227
qed "irred_imp_max_ideal";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   228
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   229
(* Pid are factorial *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   230
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   231
(* Divisor chain condition *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   232
(* proofs not finished *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   233
8707
paulson
parents: 7998
diff changeset
   234
Goal "(ALL i. I i <= I (Suc i)) ==> (n <= m & a : I n --> a : I m)";
paulson
parents: 7998
diff changeset
   235
by (induct_tac "m" 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   236
by (Blast_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   237
(* induction step *)
8707
paulson
parents: 7998
diff changeset
   238
by (rename_tac "m" 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   239
by (case_tac "n <= m" 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   240
by Auto_tac;
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   241
by (subgoal_tac "n = Suc m" 1);
8707
paulson
parents: 7998
diff changeset
   242
by (arith_tac 2);
paulson
parents: 7998
diff changeset
   243
by (Force_tac 1);
paulson
parents: 7998
diff changeset
   244
qed_spec_mp "subset_chain_lemma";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   245
8707
paulson
parents: 7998
diff changeset
   246
Goal "[| ALL i. is_ideal (I i); ALL i. I i <= I (Suc i) |] \
10833
c0844a30ea4e `` -> ` and ``` -> ``
nipkow
parents: 10789
diff changeset
   247
\     ==> is_ideal (Union (I`UNIV))";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   248
by (rtac is_idealI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   249
by Auto_tac;
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   250
by (res_inst_tac [("x", "max x xa")] exI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   251
by (rtac is_ideal_add 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   252
by (Asm_simp_tac 1);
8707
paulson
parents: 7998
diff changeset
   253
by (rtac subset_chain_lemma 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   254
by (assume_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   255
by (rtac conjI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   256
by (assume_tac 2);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   257
by (arith_tac 1);
8707
paulson
parents: 7998
diff changeset
   258
by (rtac subset_chain_lemma 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   259
by (assume_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   260
by (rtac conjI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   261
by (assume_tac 2);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   262
by (arith_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   263
by (res_inst_tac [("x", "x")] exI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   264
by (Asm_simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   265
by (res_inst_tac [("x", "x")] exI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   266
by (Asm_simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   267
qed "chain_is_ideal";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   268
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   269
(*
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   270
Goal "ALL i. ideal_of {a i} < ideal_of {a (Suc i)} ==> \
10833
c0844a30ea4e `` -> ` and ``` -> ``
nipkow
parents: 10789
diff changeset
   271
\   EX n. Union ((ideal_of o (%a. {a}))`UNIV) = ideal_of {a n}";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   272
10789
260fa2c67e3e Changed priority of dvd from 70 to 50 as befits a relation.
nipkow
parents: 10230
diff changeset
   273
Goal "wf {(a::'a::pid, b). a dvd b & ~ b dvd a}";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   274
by (simp_tac (simpset()
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   275
  addsimps [psubideal_is_dvd RS sym, wf_iff_no_infinite_down_chain]
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   276
  delsimps [psubideal_is_dvd]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   277
*)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   278
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   279
(* Primeness condition *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   280
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
   281
Goalw [thm "prime_def"] "irred a ==> prime (a::'a::pid)";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   282
by (rtac conjI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   283
by (rtac conjI 2);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   284
by (Clarify_tac 3);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
   285
by (dres_inst_tac [("I", "ideal_of {a, b}"), ("x", "1")]
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   286
  irred_imp_max_ideal 3);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   287
by (auto_tac (claset() addIs [ideal_of_is_ideal, pid_ax],
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
   288
  simpset() addsimps [thm "irred_def", not_dvd_psubideal, pid_ax]));
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   289
by (full_simp_tac (simpset() addsimps [ideal_of_2_structure]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   290
by (Clarify_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   291
by (dres_inst_tac [("f", "op* aa")] arg_cong 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   292
by (full_simp_tac (simpset() addsimps [r_distr]) 1);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
   293
by (etac subst 1);
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
   294
by (asm_simp_tac (simpset() addsimps [m_assoc RS sym]
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
   295
  delsimprocs [ring_simproc]) 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   296
qed "pid_irred_imp_prime";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   297
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   298
(* Fields are Pid *)
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   299
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10833
diff changeset
   300
Goal "a ~= 0 ==> ideal_of {a::'a::field} = UNIV";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   301
by (rtac subset_antisym 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   302
by (Simp_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   303
by (rtac subset_trans 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   304
by (rtac dvd_imp_subideal 2);
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 13654
diff changeset
   305
by (rtac (thm "field_ax") 2);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   306
by (assume_tac 2);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   307
by (simp_tac (simpset() addsimps [ideal_of_one_eq]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   308
qed "field_pideal_univ";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   309
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10833
diff changeset
   310
Goal "[| is_ideal I; I ~= {0} |] ==> {0} < I";
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   311
by (asm_simp_tac (simpset() addsimps [psubset_eq, not_sym, is_ideal_zero]) 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   312
qed "proper_ideal";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   313
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   314
Goalw [is_pideal_def] "is_ideal (I::('a::field) set) ==> is_pideal I";
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10833
diff changeset
   315
by (case_tac "I = {0}" 1);
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10833
diff changeset
   316
by (res_inst_tac [("x", "0")] exI 1);
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   317
by (asm_simp_tac (simpset() addsimps [ideal_of_zero_eq]) 1);
11093
62c2e0af1d30 Changes to HOL/Algebra:
ballarin
parents: 10833
diff changeset
   318
(* case "I ~= {0}" *)
7998
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   319
by (ftac proper_ideal 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   320
by (assume_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   321
by (dtac psubset_imp_ex_mem 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   322
by (etac exE 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   323
by (res_inst_tac [("x", "b")] exI 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   324
by (cut_inst_tac [("a", "b")] element_generates_subideal 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   325
  by (assume_tac 1); by (Blast_tac 1);
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   326
by (auto_tac (claset(), simpset() addsimps [field_pideal_univ]));
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   327
qed "field_pid";
3d0c34795831 Algebra and Polynomial theories, by Clemens Ballarin
paulson
parents:
diff changeset
   328