Quot/Examples/IntEx.thy
changeset 787 5cf83fa5b36c
parent 768 e9e205b904e2
child 833 129caba33c03
equal deleted inserted replaced
786:d6407afb913c 787:5cf83fa5b36c
     5 fun
     5 fun
     6   intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" (infix "\<approx>" 50)
     6   intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" (infix "\<approx>" 50)
     7 where
     7 where
     8   "intrel (x, y) (u, v) = (x + v = u + y)"
     8   "intrel (x, y) (u, v) = (x + v = u + y)"
     9 
     9 
    10 quotient_type my_int = "nat \<times> nat" / intrel
    10 quotient_type 
       
    11   my_int = "nat \<times> nat" / intrel
    11   apply(unfold equivp_def)
    12   apply(unfold equivp_def)
    12   apply(auto simp add: mem_def expand_fun_eq)
    13   apply(auto simp add: mem_def expand_fun_eq)
    13   done
    14   done
    14 
    15 
    15 thm quot_equiv
    16 thm quot_equiv