Quot/Examples/LarryInt.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 16 Dec 2009 14:28:48 +0100
changeset 757 c129354f2ff6
parent 753 544b05e03ec0
child 766 df053507edba
permissions -rw-r--r--
complete fix for IsaMakefile
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
header{*The Integers as Equivalence Classes over Pairs of Natural Numbers*}
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
theory LarryInt 
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
imports Nat "../QuotMain"
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
begin
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
fun
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
  intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" 
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
where
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
  "intrel (x, y) (u, v) = (x + v = u + y)"
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
quotient int = "nat \<times> nat" / intrel
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
  by (auto simp add: equivp_def expand_fun_eq)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
instantiation int :: "{zero, one, plus, uminus, minus, times, ord}"
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
begin
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
quotient_def 
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
  Zero_int_def: "0::int" as "(0::nat, 0::nat)"
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
quotient_def
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
  One_int_def: "1::int" as "(1::nat, 0::nat)"
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    25
definition
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    26
  "add_raw \<equiv> \<lambda>(x, y) (u, v). (x + (u::nat), y + (v::nat))"
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    27
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
quotient_def
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
  "(op +) :: int \<Rightarrow> int \<Rightarrow> int" 
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
as 
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    31
  "add_raw"
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    32
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    33
definition
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    34
  "uminus_raw \<equiv> \<lambda>(x::nat, y::nat). (y, x)"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
quotient_def
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
  "uminus :: int \<Rightarrow> int" 
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
as 
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    39
  "uminus_raw"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
fun
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    42
  mult_raw::"nat \<times> nat \<Rightarrow> nat \<times> nat \<Rightarrow> nat \<times> nat"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
where
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    44
  "mult_raw (x, y) (u, v) = (x*u + y*v, x*v + y*u)"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
quotient_def
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
  "(op *) :: int \<Rightarrow> int \<Rightarrow> int" 
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
as 
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    49
  "mult_raw"
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    50
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    51
definition
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    52
  "le_raw \<equiv> \<lambda>(x, y) (u, v). (x+v \<le> u+(y::nat))"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
quotient_def 
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
  le_int_def: "(op \<le>) :: int \<Rightarrow> int \<Rightarrow> bool" 
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
as 
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    57
  "le_raw"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
definition
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
  less_int_def: "z < (w::int) \<equiv> (z \<le> w & z \<noteq> w)"
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
definition
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
  diff_int_def:  "z - (w::int) \<equiv> z + (-w)"
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
instance ..
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
end
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
subsection{*Construction of the Integers*}
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    71
lemma zminus_zminus_raw:
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    72
  "uminus_raw (uminus_raw z) = z"
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    73
  by (cases z) (simp add: uminus_raw_def)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
lemma [quot_respect]:
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    76
  shows "(intrel ===> intrel) uminus_raw uminus_raw"
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    77
  by (simp add: uminus_raw_def)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
  
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
lemma zminus_zminus: 
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
  shows "- (- z) = (z::int)"
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    81
apply(lifting zminus_zminus_raw)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    84
lemma zminus_0_raw:
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    85
  shows "uminus_raw (0, 0) = (0, 0::nat)"
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    86
by (simp add: uminus_raw_def)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
lemma zminus_0: "- 0 = (0::int)"
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    89
apply(lifting zminus_0_raw)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
subsection{*Integer Addition*}
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    94
lemma zminus_zadd_distrib_raw:
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    95
  shows "uminus_raw (add_raw z w) = add_raw (uminus_raw z) (uminus_raw w)"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
by (cases z, cases w)
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
    97
   (auto simp add: add_raw_def uminus_raw_def)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
lemma [quot_respect]:
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   100
  shows "(intrel ===> intrel ===> intrel) add_raw add_raw"
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   101
by (simp add: add_raw_def)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
lemma zminus_zadd_distrib: 
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
  shows "- (z + w) = (- z) + (- w::int)"
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   105
apply(lifting zminus_zadd_distrib_raw)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   108
lemma zadd_commute_raw:
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   109
  shows "add_raw z w = add_raw w z"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
by (cases z, cases w)
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   111
   (simp add: add_raw_def)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   113
lemma zadd_commute:
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
  shows "(z::int) + w = w + z"
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   115
apply(lifting zadd_commute_raw)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   118
lemma zadd_assoc_raw:
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   119
  shows "add_raw (add_raw z1 z2) z3 = add_raw z1 (add_raw z2 z3)"
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   120
by (cases z1, cases z2, cases z3) (simp add: add_raw_def)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
lemma zadd_assoc: "((z1::int) + z2) + z3 = z1 + (z2 + z3)"
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   123
apply(lifting zadd_assoc_raw)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   126
lemma zadd_0_raw:
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
  fixes z::"nat \<times> nat"
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   128
  shows "add_raw (0, 0) z = z"
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   129
by (simp add: add_raw_def)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
(*also for the instance declaration int :: plus_ac0*)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
lemma zadd_0: "(0::int) + z = z"
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   134
apply(lifting zadd_0_raw)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   137
lemma zadd_zminus_inverse_raw:
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   138
  shows "intrel (add_raw (uminus_raw z) z) (0, 0)"
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   139
by (cases z) (simp add: add_raw_def uminus_raw_def)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
lemma zadd_zminus_inverse2: "(- z) + z = (0::int)"
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   142
apply(lifting zadd_zminus_inverse_raw)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
subsection{*Integer Multiplication*}
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   147
lemma zmult_zminus_raw:
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   148
  shows "mult_raw (uminus_raw z) w = uminus_raw (mult_raw z w)"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
apply(cases z, cases w)
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   150
apply(simp add:uminus_raw_def)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   153
lemma mult_raw_fst:
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
  assumes a: "intrel x z"
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   155
  shows "intrel (mult_raw x y) (mult_raw z y)"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
using a
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
apply(cases x, cases y, cases z)
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   158
apply(auto simp add: mult_raw.simps intrel.simps)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
apply(rename_tac u v w x y z)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
apply(subgoal_tac "u*w + z*w = y*w + v*w  &  u*x + z*x = y*x + v*x")
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
apply(simp add: mult_ac)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
apply(simp add: add_mult_distrib [symmetric])
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   165
lemma mult_raw_snd:
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
  assumes a: "intrel x z"
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   167
  shows "intrel (mult_raw y x) (mult_raw y z)"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
using a
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
apply(cases x, cases y, cases z)
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   170
apply(auto simp add: mult_raw.simps intrel.simps)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
apply(rename_tac u v w x y z)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
apply(subgoal_tac "u*w + z*w = y*w + v*w  &  u*x + z*x = y*x + v*x")
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
apply(simp add: mult_ac)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
apply(simp add: add_mult_distrib [symmetric])
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
lemma [quot_respect]:
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   178
  shows "(intrel ===> intrel ===> intrel) mult_raw mult_raw"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
apply(simp only: fun_rel.simps)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
apply(rule allI | rule impI)+
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
apply(rule equivp_transp[OF int_equivp])
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   182
apply(rule mult_raw_fst)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
apply(assumption)
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   184
apply(rule mult_raw_snd)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
apply(assumption)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
lemma zmult_zminus: "(- z) * w = - (z * (w::int))"
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   189
apply(lifting zmult_zminus_raw)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   192
lemma zmult_commute_raw: 
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   193
  shows "mult_raw z w = mult_raw w z"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
apply(cases z, cases w)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
apply(simp add: add_ac mult_ac)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
lemma zmult_commute: "(z::int) * w = w * z"
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   199
by (lifting zmult_commute_raw)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   201
lemma zmult_assoc_raw:
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   202
  shows "mult_raw (mult_raw z1 z2) z3 = mult_raw z1 (mult_raw z2 z3)"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
apply(cases z1, cases z2, cases z3)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
apply(simp add: add_mult_distrib2 mult_ac)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
lemma zmult_assoc: "((z1::int) * z2) * z3 = z1 * (z2 * z3)"
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   208
by (lifting zmult_assoc_raw)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   210
lemma zadd_mult_distrib_raw:
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   211
  shows "mult_raw (add_raw z1 z2) w = add_raw (mult_raw z1 w) (mult_raw z2 w)"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
apply(cases z1, cases z2, cases w)
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   213
apply(simp add: add_mult_distrib2 mult_ac add_raw_def)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
lemma zadd_zmult_distrib: "((z1::int) + z2) * w = (z1 * w) + (z2 * w)"
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   217
apply(lifting zadd_mult_distrib_raw)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
lemma zadd_zmult_distrib2: "(w::int) * (z1 + z2) = (w * z1) + (w * z2)"
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
by (simp add: zmult_commute [of w] zadd_zmult_distrib)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
lemma zdiff_zmult_distrib: "((z1::int) - z2) * w = (z1 * w) - (z2 * w)"
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
by (simp add: diff_int_def zadd_zmult_distrib zmult_zminus)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
lemma zdiff_zmult_distrib2: "(w::int) * (z1 - z2) = (w * z1) - (w * z2)"
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
by (simp add: zmult_commute [of w] zdiff_zmult_distrib)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
lemmas int_distrib =
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
  zadd_zmult_distrib zadd_zmult_distrib2
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
  zdiff_zmult_distrib zdiff_zmult_distrib2
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   233
lemma zmult_1_raw:
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   234
  shows "mult_raw (1, 0) z = z"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
apply(cases z)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
apply(auto)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
lemma zmult_1: "(1::int) * z = z"
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   240
apply(lifting zmult_1_raw)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
lemma zmult_1_right: "z * (1::int) = z"
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
by (rule trans [OF zmult_commute zmult_1])
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
lemma zero_not_one:
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   247
  shows "\<not>(intrel (0, 0) (1::nat, 0::nat))"
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   248
by auto
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
text{*The Integers Form A Ring*}
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
instance int :: comm_ring_1
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
proof
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
  fix i j k :: int
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
  show "(i + j) + k = i + (j + k)" by (simp add: zadd_assoc)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
  show "i + j = j + i" by (simp add: zadd_commute)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
  show "0 + i = i" by (rule zadd_0)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
  show "- i + i = 0" by (rule zadd_zminus_inverse2)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
  show "i - j = i + (-j)" by (simp add: diff_int_def)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
  show "(i * j) * k = i * (j * k)" by (rule zmult_assoc)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
  show "i * j = j * i" by (rule zmult_commute)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
  show "1 * i = i" by (rule zmult_1) 
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
  show "(i + j) * k = i * k + j * k" by (simp add: int_distrib)
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   263
  show "0 \<noteq> (1::int)" by (lifting zero_not_one)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
qed
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
subsection{*The @{text "\<le>"} Ordering*}
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   269
lemma zle_refl_raw:
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   270
  "le_raw w w"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
apply(cases w)
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   272
apply(simp add: le_raw_def)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
lemma [quot_respect]:
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   276
  shows "(intrel ===> intrel ===> op =) le_raw le_raw"
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   277
by (auto) (simp_all add: le_raw_def)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
lemma zle_refl: "w \<le> (w::int)"
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   280
apply(lifting zle_refl_raw)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   283
lemma zle_trans_raw:
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   284
  shows "\<lbrakk>le_raw i j; le_raw j k\<rbrakk> \<Longrightarrow> le_raw i k"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
apply(cases i, cases j, cases k)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   286
apply(auto)
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   287
apply(simp add:le_raw_def)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   289
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
lemma zle_trans: "\<lbrakk>i \<le> j; j \<le> k\<rbrakk> \<Longrightarrow> i \<le> (k::int)"
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   291
apply(lifting zle_trans_raw)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   292
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   293
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   294
lemma zle_anti_sym_raw:
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   295
  shows "\<lbrakk>le_raw z w; le_raw w z\<rbrakk> \<Longrightarrow> intrel z w"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   296
apply(cases z, cases w)
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   297
apply(auto iff: le_raw_def)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   298
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   299
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   300
lemma zle_anti_sym: "\<lbrakk>z \<le> w; w \<le> z\<rbrakk> \<Longrightarrow> z = (w::int)"
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   301
apply(lifting zle_anti_sym_raw)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   302
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   303
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   304
(* Axiom 'order_less_le' of class 'order': *)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
lemma zless_le: "((w::int) < z) = (w \<le> z & w \<noteq> z)"
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
by (simp add: less_int_def)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   308
instance int :: order
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   309
apply(intro_classes)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   310
apply(auto intro: zle_refl zle_trans zle_anti_sym zless_le simp add: less_int_def)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   312
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   313
(* Axiom 'linorder_linear' of class 'linorder': *)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   314
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   315
lemma zle_linear_raw:
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   316
  "le_raw z w \<or> le_raw w z"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   317
apply(cases w, cases z)
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   318
apply(auto iff: le_raw_def)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   319
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   320
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   322
lemma zle_linear: "(z::int) \<le> w \<or> w \<le> z"
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   323
apply(lifting zle_linear_raw)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   324
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   325
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   326
instance int :: linorder
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
proof qed (rule zle_linear)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   329
lemma zadd_left_mono_raw:
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   330
  shows "le_raw i j \<Longrightarrow> le_raw (add_raw k i) (add_raw k j)"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
apply(cases k)
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   332
apply(auto simp add: add_raw_def le_raw_def)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   334
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
lemma zadd_left_mono: "i \<le> j \<Longrightarrow> k + i \<le> k + (j::int)"
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   336
apply(lifting zadd_left_mono_raw)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   341
subsection{*Magnitide of an Integer, as a Natural Number: @{term nat}*}
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   342
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   343
definition
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   344
  "nat_raw \<equiv> \<lambda>(x, y).x - (y::nat)"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   345
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
quotient_def
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   347
  "nat2::int\<Rightarrow>nat"
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
as
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   349
  "nat_raw"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   350
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   351
abbreviation
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   352
  "less_raw x y \<equiv> (le_raw x y \<and> \<not>(x = y))"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   353
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   354
lemma nat_le_eq_zle_raw:
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   355
  shows "less_raw (0, 0) w \<or> le_raw (0, 0) z \<Longrightarrow> (nat_raw w \<le> nat_raw z) = (le_raw w z)"
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   356
apply(auto)
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   357
sorry
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   358
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   359
lemma [quot_respect]:
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   360
  shows "(intrel ===> op =) nat_raw nat_raw"
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   361
apply(auto iff: nat_raw_def)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   363
717
337dd914e1cb deleted struct_match by Pattern.match (fixes a problem in LarryInt)
Christian Urban <urbanc@in.tum.de>
parents: 715
diff changeset
   364
ML {*
337dd914e1cb deleted struct_match by Pattern.match (fixes a problem in LarryInt)
Christian Urban <urbanc@in.tum.de>
parents: 715
diff changeset
   365
  let
337dd914e1cb deleted struct_match by Pattern.match (fixes a problem in LarryInt)
Christian Urban <urbanc@in.tum.de>
parents: 715
diff changeset
   366
   val parser = Args.context -- Scan.lift Args.name_source
337dd914e1cb deleted struct_match by Pattern.match (fixes a problem in LarryInt)
Christian Urban <urbanc@in.tum.de>
parents: 715
diff changeset
   367
   fun term_pat (ctxt, str) =
337dd914e1cb deleted struct_match by Pattern.match (fixes a problem in LarryInt)
Christian Urban <urbanc@in.tum.de>
parents: 715
diff changeset
   368
      str |> ProofContext.read_term_pattern ctxt
337dd914e1cb deleted struct_match by Pattern.match (fixes a problem in LarryInt)
Christian Urban <urbanc@in.tum.de>
parents: 715
diff changeset
   369
          |> ML_Syntax.print_term
337dd914e1cb deleted struct_match by Pattern.match (fixes a problem in LarryInt)
Christian Urban <urbanc@in.tum.de>
parents: 715
diff changeset
   370
          |> ML_Syntax.atomic
337dd914e1cb deleted struct_match by Pattern.match (fixes a problem in LarryInt)
Christian Urban <urbanc@in.tum.de>
parents: 715
diff changeset
   371
in
337dd914e1cb deleted struct_match by Pattern.match (fixes a problem in LarryInt)
Christian Urban <urbanc@in.tum.de>
parents: 715
diff changeset
   372
   ML_Antiquote.inline "term_pat" (parser >> term_pat)
337dd914e1cb deleted struct_match by Pattern.match (fixes a problem in LarryInt)
Christian Urban <urbanc@in.tum.de>
parents: 715
diff changeset
   373
end
337dd914e1cb deleted struct_match by Pattern.match (fixes a problem in LarryInt)
Christian Urban <urbanc@in.tum.de>
parents: 715
diff changeset
   374
337dd914e1cb deleted struct_match by Pattern.match (fixes a problem in LarryInt)
Christian Urban <urbanc@in.tum.de>
parents: 715
diff changeset
   375
*}
337dd914e1cb deleted struct_match by Pattern.match (fixes a problem in LarryInt)
Christian Urban <urbanc@in.tum.de>
parents: 715
diff changeset
   376
337dd914e1cb deleted struct_match by Pattern.match (fixes a problem in LarryInt)
Christian Urban <urbanc@in.tum.de>
parents: 715
diff changeset
   377
consts test::"'b \<Rightarrow> 'b \<Rightarrow> 'b"
337dd914e1cb deleted struct_match by Pattern.match (fixes a problem in LarryInt)
Christian Urban <urbanc@in.tum.de>
parents: 715
diff changeset
   378
337dd914e1cb deleted struct_match by Pattern.match (fixes a problem in LarryInt)
Christian Urban <urbanc@in.tum.de>
parents: 715
diff changeset
   379
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   380
lemma nat_le_eq_zle: "0 < w \<or> 0 \<le> z \<Longrightarrow> (nat2 w \<le> nat2 z) = (w\<le>z)"
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   381
unfolding less_int_def
753
544b05e03ec0 Removed lambdas on the right hand side. This fixes all 'PROBLEM' comments.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 717
diff changeset
   382
apply(lifting nat_le_eq_zle_raw)
715
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   383
done
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   384
3d7a9d4d2bb6 added Int example from Larry
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   385
end