Quot/Examples/LarryInt.thy
changeset 1139 c4001cda9da3
parent 1128 17ca92ab4660
child 1153 2ad8e66de294
equal deleted inserted replaced
1138:93c9022ba5e9 1139:c4001cda9da3
    14   by (auto simp add: equivp_def expand_fun_eq)
    14   by (auto simp add: equivp_def expand_fun_eq)
    15 
    15 
    16 instantiation int :: "{zero, one, plus, uminus, minus, times, ord}"
    16 instantiation int :: "{zero, one, plus, uminus, minus, times, ord}"
    17 begin
    17 begin
    18 
    18 
    19 quotient_definition 
    19 quotient_definition
    20   Zero_int_def: "0::int" as "(0::nat, 0::nat)"
    20   Zero_int_def: "0::int" is "(0::nat, 0::nat)"
    21 
    21 
    22 quotient_definition
    22 quotient_definition
    23   One_int_def: "1::int" as "(1::nat, 0::nat)"
    23   One_int_def: "1::int" is "(1::nat, 0::nat)"
    24 
    24 
    25 definition
    25 definition
    26   "add_raw \<equiv> \<lambda>(x, y) (u, v). (x + (u::nat), y + (v::nat))"
    26   "add_raw \<equiv> \<lambda>(x, y) (u, v). (x + (u::nat), y + (v::nat))"
    27 
    27 
    28 quotient_definition
    28 quotient_definition
    29   "(op +) :: int \<Rightarrow> int \<Rightarrow> int" 
    29   "(op +) :: int \<Rightarrow> int \<Rightarrow> int" 
    30 as 
    30 is
    31   "add_raw"
    31   "add_raw"
    32 
    32 
    33 definition
    33 definition
    34   "uminus_raw \<equiv> \<lambda>(x::nat, y::nat). (y, x)"
    34   "uminus_raw \<equiv> \<lambda>(x::nat, y::nat). (y, x)"
    35 
    35 
    36 quotient_definition
    36 quotient_definition
    37   "uminus :: int \<Rightarrow> int" 
    37   "uminus :: int \<Rightarrow> int" 
    38 as 
    38 is
    39   "uminus_raw"
    39   "uminus_raw"
    40 
    40 
    41 fun
    41 fun
    42   mult_raw::"nat \<times> nat \<Rightarrow> nat \<times> nat \<Rightarrow> nat \<times> nat"
    42   mult_raw::"nat \<times> nat \<Rightarrow> nat \<times> nat \<Rightarrow> nat \<times> nat"
    43 where
    43 where
    44   "mult_raw (x, y) (u, v) = (x*u + y*v, x*v + y*u)"
    44   "mult_raw (x, y) (u, v) = (x*u + y*v, x*v + y*u)"
    45 
    45 
    46 quotient_definition
    46 quotient_definition
    47   "(op *) :: int \<Rightarrow> int \<Rightarrow> int" 
    47   "(op *) :: int \<Rightarrow> int \<Rightarrow> int" 
    48 as 
    48 is
    49   "mult_raw"
    49   "mult_raw"
    50 
    50 
    51 definition
    51 definition
    52   "le_raw \<equiv> \<lambda>(x, y) (u, v). (x+v \<le> u+(y::nat))"
    52   "le_raw \<equiv> \<lambda>(x, y) (u, v). (x+v \<le> u+(y::nat))"
    53 
    53 
    54 quotient_definition 
    54 quotient_definition 
    55   le_int_def: "(op \<le>) :: int \<Rightarrow> int \<Rightarrow> bool" 
    55   le_int_def: "(op \<le>) :: int \<Rightarrow> int \<Rightarrow> bool" 
    56 as 
    56 is
    57   "le_raw"
    57   "le_raw"
    58 
    58 
    59 definition
    59 definition
    60   less_int_def: "z < (w::int) \<equiv> (z \<le> w & z \<noteq> w)"
    60   less_int_def: "z < (w::int) \<equiv> (z \<le> w & z \<noteq> w)"
    61 
    61 
   367 definition
   367 definition
   368   "nat_raw \<equiv> \<lambda>(x, y).x - (y::nat)"
   368   "nat_raw \<equiv> \<lambda>(x, y).x - (y::nat)"
   369 
   369 
   370 quotient_definition
   370 quotient_definition
   371   "nat2::int \<Rightarrow> nat"
   371   "nat2::int \<Rightarrow> nat"
   372 as
   372 is
   373   "nat_raw"
   373   "nat_raw"
   374 
   374 
   375 abbreviation
   375 abbreviation
   376   "less_raw x y \<equiv> (le_raw x y \<and> \<not>(x = y))"
   376   "less_raw x y \<equiv> (le_raw x y \<and> \<not>(x = y))"
   377 
   377