Quot/Examples/AbsRepTest.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 24 Dec 2009 22:28:19 +0100
changeset 787 5cf83fa5b36c
parent 786 d6407afb913c
child 790 3a48ffcf0f9a
permissions -rw-r--r--
made the quotient_type definition more like typedef; now type variables need to be explicitly given

theory AbsRepTest
imports "../QuotMain" "../QuotList" "../QuotOption" "../QuotSum" "../QuotProd" List
begin

print_maps

quotient_type 
  'a fset = "'a list" / "\<lambda>xs ys. \<forall>e. e \<in> set xs \<longleftrightarrow> e \<in> set ys" 
  apply(rule equivpI)
  unfolding reflp_def symp_def transp_def
  apply(auto)
  done

quotient_type 
  'a foo = "('a * 'a) list" / "\<lambda>(xs::('a * 'a) list) ys. \<forall>e. e \<in> set xs \<longleftrightarrow> e \<in> set ys" 
  apply(rule equivpI)
  unfolding reflp_def symp_def transp_def
  apply(auto)
  done

print_quotients

ML{*
Quotient_Info.quotient_rules_get @{context}
*}

quotient_type int = "nat \<times> nat" / "\<lambda>(x, y) (u, v). x + v = u + (y::nat)"
  apply(rule equivpI)
  unfolding reflp_def symp_def transp_def
  apply(auto)
  done

print_quotients

ML {*
open Quotient_Term;
*}

ML {*
absrep_fun_chk absF @{context} 
     (@{typ "(('a * 'a) list) list"}, 
      @{typ "(('a * 'a) fset) fset"})
|> Syntax.string_of_term @{context}
|> writeln
*}

(*
ML {*
absrep_fun_chk absF @{context} 
     (@{typ "('a * 'a) list"}, 
      @{typ "'a foo"})
|> Syntax.string_of_term @{context}
|> writeln
*}
*)

term "option_map (prod_fun (abs_fset \<circ> map abs_int) id)"
term "option_map (prod_fun (map rep_int \<circ> rep_fset) id)"
term "option_map (map rep_int \<circ> rep_fset)"
term "option_map (abs_fset o (map abs_int))"
term "abs_fset"

ML {*
absrep_fun_chk absF @{context} (@{typ "(nat * nat) list"}, @{typ "int fset"})
|> Syntax.string_of_term @{context}
|> writeln
*}

term "map abs_int"
term "abs_fset o map abs_int"


ML {*
absrep_fun_chk absF @{context} (@{typ "('a list) list"}, @{typ "('a fset) fset"})
|> Syntax.string_of_term @{context}
|> writeln
*}

ML {*
absrep_fun_chk absF @{context} (@{typ "('a list) list \<Rightarrow> 'a"}, @{typ "('a fset) fset \<Rightarrow> 'a"})
|> Syntax.string_of_term @{context}
|> writeln
*}

end