src/HOL/ex/LocaleGroup.ML
author paulson
Fri, 11 Dec 1998 10:41:53 +0100
changeset 6024 cb87f103d114
parent 5848 99dea3c24efb
permissions -rw-r--r--
new Close_locale synatx
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/ex/LocaleGroup.ML
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
     3
    Author:     Florian Kammueller, University of Cambridge
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
     4
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
     5
Group theory via records and locales.
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
     6
*)
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
     7
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
     8
Open_locale "groups";
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
     9
print_locales LocaleGroup.thy;
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    10
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    11
val simp_G = simplify (simpset() addsimps [Group_def]) (thm "Group_G");
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    12
Addsimps [simp_G, thm "Group_G"];
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    13
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    14
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    15
Goal "e : carrier G";
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    16
by (simp_tac (simpset() addsimps [thm "e_def"]) 1);
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    17
qed "e_closed";
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    18
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    19
(* Mit dieser Def ist es halt schwierig *)
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    20
Goal "op # : carrier G -> carrier G -> carrier G";
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    21
by (res_inst_tac [("t","op #")] ssubst 1);
5318
72bf8039b53f expandshort
paulson
parents: 5250
diff changeset
    22
by (rtac ext 1);
72bf8039b53f expandshort
paulson
parents: 5250
diff changeset
    23
by (rtac ext 1);
72bf8039b53f expandshort
paulson
parents: 5250
diff changeset
    24
by (rtac meta_eq_to_obj_eq 1);
72bf8039b53f expandshort
paulson
parents: 5250
diff changeset
    25
by (rtac (thm "binop_def") 1);
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    26
by (Asm_full_simp_tac 1);
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    27
qed "binop_funcset";
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    28
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    29
Goal "[| x: carrier G; y: carrier G |] ==> x # y : carrier G";
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    30
by (asm_simp_tac
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    31
    (simpset() addsimps [binop_funcset RS funcset_mem RS funcset_mem]) 1);
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    32
qed "binop_closed";
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    33
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    34
Addsimps [binop_closed, e_closed];
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    35
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    36
Goal "INV : carrier G -> carrier G";
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    37
by (asm_simp_tac (simpset() addsimps [thm "inv_def"]) 1);
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    38
qed "inv_funcset";
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    39
5848
99dea3c24efb changed inverse syntax from x-| to i(x)
paulson
parents: 5845
diff changeset
    40
Goal "x: carrier G ==> i(x) : carrier G";
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    41
by (asm_simp_tac (simpset() addsimps [inv_funcset RS funcset_mem]) 1);
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    42
qed "inv_closed"; 
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    43
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    44
Goal "x: carrier G ==> e # x = x";
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    45
by (asm_simp_tac (simpset() addsimps [thm "e_def", thm "binop_def"]) 1);
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    46
qed "e_ax1";
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    47
5848
99dea3c24efb changed inverse syntax from x-| to i(x)
paulson
parents: 5845
diff changeset
    48
Goal "x: carrier G ==> i(x) # x = e";
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    49
by (asm_simp_tac
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    50
    (simpset() addsimps [thm "binop_def", thm "inv_def", thm "e_def"]) 1);
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    51
qed "inv_ax2";
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    52
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    53
Addsimps [inv_closed, e_ax1, inv_ax2];
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    54
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    55
Goal "[| x: carrier G; y: carrier G; z: carrier G |]\
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    56
\               ==> (x # y) # z = x # (y # z)";
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    57
by (asm_simp_tac (simpset() addsimps [thm "binop_def"]) 1);
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    58
qed "binop_assoc";
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    59
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    60
Goal "[|f : A -> A -> A; i: A -> A; e1: A;\
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    61
\        ! x: A. (f (i x) x = e1); ! x: A. (f e1 x = x);\
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    62
\        ! x: A. ! y: A. ! z: A.(f (f x y) z = f (x) (f y z)) |] \
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    63
\     ==> (| carrier = A, bin_op = f, inverse = i, unit = e1 |) : Group";
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    64
by (asm_simp_tac (simpset() addsimps [Group_def]) 1);
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    65
qed "GroupI";
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    66
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    67
(*****)
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    68
(* Now the real derivations *)
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    69
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    70
Goal "[| x # y  =  x # z;  \
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    71
\        x : carrier G ; y : carrier G; z : carrier G |] ==> y = z";
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    72
by (res_inst_tac [("P","%r. r = z")] (e_ax1 RS subst) 1);
5318
72bf8039b53f expandshort
paulson
parents: 5250
diff changeset
    73
by (assume_tac 1);
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    74
(* great: we can use the nice syntax even in res_inst_tac *)
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    75
by (res_inst_tac [("P","%r. r # y = z")] (inv_ax2 RS subst) 1);
5318
72bf8039b53f expandshort
paulson
parents: 5250
diff changeset
    76
by (assume_tac 1);
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    77
by (asm_simp_tac (simpset() delsimps [inv_ax2] addsimps [binop_assoc]) 1);
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    78
by (asm_simp_tac (simpset() addsimps [binop_assoc RS sym]) 1);
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    79
qed "left_cancellation";
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    80
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    81
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    82
(* Here are the other directions of basic lemmas. 
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    83
   They needed a cancellation (left) to be able to show the other
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    84
   directions of inverse and unity axiom.*)
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    85
Goal "x: carrier G ==> x # e = x";
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    86
by (rtac left_cancellation 1);
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    87
by (etac inv_closed 2);
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    88
by (auto_tac (claset(), simpset() addsimps [binop_assoc RS sym]));
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    89
qed "e_ax2";
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    90
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    91
Addsimps [e_ax2];
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    92
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    93
Goal "[| x: carrier G; x # x = x |] ==> x = e";
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    94
by (forw_inst_tac [("P","%y. x # x = y")] (e_ax2 RS ssubst) 1);
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    95
by (etac left_cancellation 2);
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    96
by Auto_tac;
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
    97
qed "idempotent_e";
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
    98
5848
99dea3c24efb changed inverse syntax from x-| to i(x)
paulson
parents: 5845
diff changeset
    99
Goal  "x: carrier G ==> x # i(x) = e";
5318
72bf8039b53f expandshort
paulson
parents: 5250
diff changeset
   100
by (rtac idempotent_e 1);
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   101
by (Asm_simp_tac 1);
5848
99dea3c24efb changed inverse syntax from x-| to i(x)
paulson
parents: 5845
diff changeset
   102
by (subgoal_tac "(x # i(x)) # x # i(x) = x # (i(x) # x) # i(x)" 1);
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   103
by (asm_simp_tac (simpset() delsimps [inv_ax2]
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   104
			    addsimps [binop_assoc]) 2);
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   105
by Auto_tac;
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   106
qed "inv_ax1";
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   107
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   108
Addsimps [inv_ax1];
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   109
5848
99dea3c24efb changed inverse syntax from x-| to i(x)
paulson
parents: 5845
diff changeset
   110
Goal "[| x # y = e; x: carrier G; y: carrier G |] ==> y = i(x)";
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   111
by (res_inst_tac [("x","x")] left_cancellation 1);
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   112
by Auto_tac;
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   113
qed "inv_unique";
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   114
5848
99dea3c24efb changed inverse syntax from x-| to i(x)
paulson
parents: 5845
diff changeset
   115
Goal "x : carrier G ==> i(i(x)) = x";
99dea3c24efb changed inverse syntax from x-| to i(x)
paulson
parents: 5845
diff changeset
   116
by (res_inst_tac [("x","i(x)")] left_cancellation 1);
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   117
by Auto_tac;
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   118
qed "inv_inv";
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   119
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   120
Addsimps [inv_inv];
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   121
5848
99dea3c24efb changed inverse syntax from x-| to i(x)
paulson
parents: 5845
diff changeset
   122
Goal "[| x : carrier G; y : carrier G |] ==> i(x # y) = i(y) # i(x)";
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   123
by (rtac (inv_unique RS sym) 1);
5848
99dea3c24efb changed inverse syntax from x-| to i(x)
paulson
parents: 5845
diff changeset
   124
by (subgoal_tac "(x # y) # i(y) # i(x) = x # (y # i(y)) # i(x)" 1);
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   125
by (asm_simp_tac (simpset() delsimps [inv_ax1, inv_ax2]
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   126
			    addsimps [binop_assoc]) 2);
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   127
by Auto_tac;
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   128
qed "inv_prod";
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
   129
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
   130
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   131
Goal "[| y # x = z # x;  x : carrier G; y : carrier G; \
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   132
\        z : carrier G |] ==> y = z";
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
   133
by (res_inst_tac [("P","%r. r = z")] (e_ax2 RS subst) 1);
5318
72bf8039b53f expandshort
paulson
parents: 5250
diff changeset
   134
by (assume_tac 1);
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   135
by (res_inst_tac [("P","%r. y # r = z")] (inv_ax1 RS subst) 1);
5318
72bf8039b53f expandshort
paulson
parents: 5250
diff changeset
   136
by (assume_tac 1);
5845
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   137
by (asm_simp_tac (simpset() delsimps [inv_ax1] 
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   138
		  addsimps [binop_assoc RS sym]) 1);
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   139
by (asm_simp_tac (simpset() addsimps [binop_assoc]) 1);
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   140
qed "right_cancellation";
eb183b062eae Big simplification of proofs.
paulson
parents: 5318
diff changeset
   141
6024
cb87f103d114 new Close_locale synatx
paulson
parents: 5848
diff changeset
   142
Close_locale "groups";
5250
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
   143
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
   144
(* example what happens if export *)
1bff4b1e5ba9 added LocaleGroup, PiSets examples;
wenzelm
parents:
diff changeset
   145
val Left_cancellation = export left_cancellation;