author | wenzelm |
Sun, 24 Feb 2019 12:49:32 +0100 | |
changeset 69838 | 4419d4d675c3 |
parent 68035 | 6d7cc6723978 |
permissions | -rw-r--r-- |
68035
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
1 |
(* Author: Florian Haftmann, TUM |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
2 |
*) |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
3 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
4 |
section \<open>Proof of concept for residue rings over int using type numerals\<close> |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
5 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
6 |
theory Residue_Ring |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
7 |
imports Main "HOL-Library.Type_Length" |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
8 |
begin |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
9 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
10 |
class len2 = len0 + |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
11 |
assumes len_ge_2 [iff]: "2 \<le> LENGTH('a)" |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
12 |
begin |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
13 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
14 |
subclass len |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
15 |
proof |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
16 |
show "0 < LENGTH('a)" using len_ge_2 |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
17 |
by arith |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
18 |
qed |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
19 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
20 |
lemma len_not_eq_Suc_0 [simp]: |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
21 |
"LENGTH('a) \<noteq> Suc 0" |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
22 |
using len_ge_2 by arith |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
23 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
24 |
end |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
25 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
26 |
instance bit0 and bit1 :: (len) len2 |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
27 |
by standard (simp_all add: Suc_le_eq) |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
28 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
29 |
quotient_type (overloaded) 'a residue_ring = int / "\<lambda>k l. k mod int (LENGTH('a)) = l mod int (LENGTH('a::len2))" |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
30 |
by (auto intro!: equivpI reflpI sympI transpI) |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
31 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
32 |
context semiring_1 |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
33 |
begin |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
34 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
35 |
lift_definition repr :: "'b::len2 residue_ring \<Rightarrow> 'a" |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
36 |
is "\<lambda>k. of_nat (nat (k mod int (LENGTH('b))))" |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
37 |
by simp |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
38 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
39 |
end |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
40 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
41 |
instantiation residue_ring :: (len2) comm_ring_1 |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
42 |
begin |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
43 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
44 |
lift_definition zero_residue_ring :: "'a residue_ring" |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
45 |
is 0 . |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
46 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
47 |
lift_definition one_residue_ring :: "'a residue_ring" |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
48 |
is 1 . |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
49 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
50 |
lift_definition plus_residue_ring :: "'a residue_ring \<Rightarrow> 'a residue_ring \<Rightarrow> 'a residue_ring" |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
51 |
is plus |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
52 |
by (fact mod_add_cong) |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
53 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
54 |
lift_definition uminus_residue_ring :: "'a residue_ring \<Rightarrow> 'a residue_ring" |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
55 |
is uminus |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
56 |
by (fact mod_minus_cong) |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
57 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
58 |
lift_definition minus_residue_ring :: "'a residue_ring \<Rightarrow> 'a residue_ring \<Rightarrow> 'a residue_ring" |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
59 |
is minus |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
60 |
by (fact mod_diff_cong) |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
61 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
62 |
lift_definition times_residue_ring :: "'a residue_ring \<Rightarrow> 'a residue_ring \<Rightarrow> 'a residue_ring" |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
63 |
is times |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
64 |
by (fact mod_mult_cong) |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
65 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
66 |
instance |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
67 |
by (standard; transfer) (simp_all add: algebra_simps mod_eq_0_iff_dvd) |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
68 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
69 |
end |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
70 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
71 |
context |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
72 |
includes lifting_syntax |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
73 |
begin |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
74 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
75 |
lemma [transfer_rule]: |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
76 |
"((\<longleftrightarrow>) ===> pcr_residue_ring) of_bool of_bool" |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
77 |
by (unfold of_bool_def [abs_def]; transfer_prover) |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
78 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
79 |
lemma [transfer_rule]: |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
80 |
"((=) ===> pcr_residue_ring) numeral numeral" |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
81 |
by (rule transfer_rule_numeral; transfer_prover) |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
82 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
83 |
lemma [transfer_rule]: |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
84 |
"((=) ===> pcr_residue_ring) int of_nat" |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
85 |
by (rule transfer_rule_of_nat; transfer_prover) |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
86 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
87 |
end |
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
88 |
|
6d7cc6723978
proof of concept for residue rings over int using type numerals
haftmann
parents:
diff
changeset
|
89 |
end |