Quot/Examples/AbsRepTest.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 23 Dec 2009 22:42:30 +0100
changeset 783 06e17083e90b
parent 780 a24e26f5488c
child 786 d6407afb913c
permissions -rw-r--r--
modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory AbsRepTest
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports "../QuotMain" "../QuotList" "../QuotOption" "../QuotSum" "../QuotProd" List
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
quotient_type 
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
  fset = "'a list" / "\<lambda>(xs::'a list) ys. \<forall>e. e \<in> set xs \<longleftrightarrow> e \<in> set ys" 
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
  apply(rule equivpI)
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
  unfolding reflp_def symp_def transp_def
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
  apply(auto)
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
  done
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
783
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 780
diff changeset
    12
print_quotients
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 780
diff changeset
    13
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 780
diff changeset
    14
ML{*
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 780
diff changeset
    15
Quotient_Info.quotient_rules_get @{context}
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 780
diff changeset
    16
*}
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 780
diff changeset
    17
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
fun
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
  intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool"
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
where
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
  "intrel (x, y) (u, v) = (x + v = u + y)"
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
quotient_type int = "nat \<times> nat" / intrel
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
  sorry
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
783
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 780
diff changeset
    26
print_quotients
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 780
diff changeset
    27
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
ML {*
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
open Quotient_Term;
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
*}
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
ML {*
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
absrep_fun_chk absF @{context} (@{typ "(((nat * nat) list) * 'a) option"}, @{typ "((int fset) * 'a) option"})
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
|> Syntax.string_of_term @{context}
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
|> writeln
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
*}
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
term "option_map (prod_fun (abs_fset \<circ> map abs_int) id)"
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
term "option_map (prod_fun (map rep_int \<circ> rep_fset) id)"
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
term "option_map (map rep_int \<circ> rep_fset)"
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
term "option_map (abs_fset o (map abs_int))"
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
term "abs_fset"
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
ML {*
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
absrep_fun_chk absF @{context} (@{typ "(nat * nat) list"}, @{typ "int fset"})
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
|> Syntax.string_of_term @{context}
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
|> writeln
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
*}
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
term "map abs_int"
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
term "abs_fset o map abs_int"
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
ML {*
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
absrep_fun_chk absF @{context} (@{typ "('a list) list"}, @{typ "('a fset) fset"})
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
|> Syntax.string_of_term @{context}
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
|> writeln
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
*}
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
ML {*
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
absrep_fun_chk absF @{context} (@{typ "('a list) list \<Rightarrow> 'a"}, @{typ "('a fset) fset \<Rightarrow> 'a"})
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
|> Syntax.string_of_term @{context}
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
|> writeln
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
*}
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
end