Quot/Examples/AbsRepTest.thy
author Christian Urban <urbanc@in.tum.de>
Sat, 09 Jan 2010 08:52:06 +0100
changeset 831 224909b9399f
parent 824 cedfe2a71298
child 852 67e5da3a356a
permissions -rw-r--r--
removed obsolete equiv_relation and rnamed new_equiv_relation
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
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
     5
ML {* open Quotient_Term *}
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
     6
795
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
     7
ML {*
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
     8
fun test_funs flag ctxt (rty, qty) =
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
     9
  (absrep_fun_chk flag ctxt (rty, qty)
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    10
   |> Syntax.string_of_term ctxt
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    11
   |> writeln;
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    12
   equiv_relation_chk ctxt (rty, qty) 
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    13
   |> Syntax.string_of_term ctxt
807
a5495a323b49 added a new version of equiv_relation (is not yet used anywhere except in AbsRepTest)
Christian Urban <urbanc@in.tum.de>
parents: 803
diff changeset
    14
   |> writeln)
795
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    15
*}
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    16
795
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    17
definition
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    18
  erel1 (infixl "\<approx>1" 50)
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    19
where
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    20
  "erel1 \<equiv> \<lambda>xs ys. \<forall>e. e \<in> set xs \<longleftrightarrow> e \<in> set ys"
787
5cf83fa5b36c made the quotient_type definition more like typedef; now type variables need to be explicitly given
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    21
5cf83fa5b36c made the quotient_type definition more like typedef; now type variables need to be explicitly given
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    22
quotient_type 
795
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    23
  'a fset = "'a list" / erel1
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
    24
  apply(rule equivpI)
795
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    25
  unfolding erel1_def reflp_def symp_def transp_def
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    26
  by auto
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    27
795
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    28
definition
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    29
  erel2 (infixl "\<approx>2" 50)
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    30
where
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    31
  "erel2 \<equiv> \<lambda>(xs::('a * 'a) list) ys. \<forall>e. e \<in> set xs \<longleftrightarrow> e \<in> set ys"
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    32
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    33
quotient_type 
795
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    34
  'a foo = "('a * 'a) list" / erel2
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    35
  apply(rule equivpI)
795
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    36
  unfolding erel2_def reflp_def symp_def transp_def
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    37
  by auto
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    38
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    39
definition
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    40
  erel3 (infixl "\<approx>3" 50)
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    41
where
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    42
  "erel3 \<equiv> \<lambda>(xs::('a * int) list) ys. \<forall>e. e \<in> set xs \<longleftrightarrow> e \<in> set ys"
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    43
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    44
quotient_type 
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    45
  'a bar = "('a * int) list" / "erel3"
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    46
  apply(rule equivpI)
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    47
  unfolding erel3_def reflp_def symp_def transp_def
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    48
  by auto
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    49
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    50
fun
795
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    51
  intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" (infixl "\<approx>4" 50)
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    52
where
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    53
  "intrel (x, y) (u, v) = (x + v = u + y)"
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    54
795
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    55
quotient_type myint = "nat \<times> nat" / intrel
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    56
  by (auto simp add: equivp_def expand_fun_eq)
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    57
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    58
ML {*
795
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    59
test_funs absF @{context} 
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    60
     (@{typ "nat \<times> nat"}, 
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    61
      @{typ "myint"})
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    62
*}
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    63
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    64
ML {*
795
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    65
test_funs absF @{context} 
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    66
     (@{typ "('a * 'a) list"}, 
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    67
      @{typ "'a foo"})
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    68
*}
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    69
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    70
ML {*
803
6f6ee78c7357 a slight change to abs/rep generation
Christian Urban <urbanc@in.tum.de>
parents: 796
diff changeset
    71
test_funs repF @{context} 
6f6ee78c7357 a slight change to abs/rep generation
Christian Urban <urbanc@in.tum.de>
parents: 796
diff changeset
    72
     (@{typ "(('a * 'a) list * 'b)"}, 
6f6ee78c7357 a slight change to abs/rep generation
Christian Urban <urbanc@in.tum.de>
parents: 796
diff changeset
    73
      @{typ "('a foo * 'b)"})
6f6ee78c7357 a slight change to abs/rep generation
Christian Urban <urbanc@in.tum.de>
parents: 796
diff changeset
    74
*}
6f6ee78c7357 a slight change to abs/rep generation
Christian Urban <urbanc@in.tum.de>
parents: 796
diff changeset
    75
6f6ee78c7357 a slight change to abs/rep generation
Christian Urban <urbanc@in.tum.de>
parents: 796
diff changeset
    76
ML {*
795
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    77
test_funs absF @{context} 
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    78
     (@{typ "(('a list) * int) list"}, 
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    79
      @{typ "('a fset) bar"})
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    80
*}
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    81
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
    82
ML {*
795
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    83
test_funs absF @{context} 
807
a5495a323b49 added a new version of equiv_relation (is not yet used anywhere except in AbsRepTest)
Christian Urban <urbanc@in.tum.de>
parents: 803
diff changeset
    84
     (@{typ "('a list)"}, 
a5495a323b49 added a new version of equiv_relation (is not yet used anywhere except in AbsRepTest)
Christian Urban <urbanc@in.tum.de>
parents: 803
diff changeset
    85
      @{typ "('a fset)"})
a5495a323b49 added a new version of equiv_relation (is not yet used anywhere except in AbsRepTest)
Christian Urban <urbanc@in.tum.de>
parents: 803
diff changeset
    86
*}
a5495a323b49 added a new version of equiv_relation (is not yet used anywhere except in AbsRepTest)
Christian Urban <urbanc@in.tum.de>
parents: 803
diff changeset
    87
a5495a323b49 added a new version of equiv_relation (is not yet used anywhere except in AbsRepTest)
Christian Urban <urbanc@in.tum.de>
parents: 803
diff changeset
    88
ML {*
a5495a323b49 added a new version of equiv_relation (is not yet used anywhere except in AbsRepTest)
Christian Urban <urbanc@in.tum.de>
parents: 803
diff changeset
    89
test_funs absF @{context} 
795
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    90
     (@{typ "('a list) list"}, 
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    91
      @{typ "('a fset) fset"})
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    92
*}
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    93
808
90bde96f5dd1 proper handling of error messages (code copy - maybe this can be avoided)
Christian Urban <urbanc@in.tum.de>
parents: 807
diff changeset
    94
811
Christian Urban <urbanc@in.tum.de>
parents: 810
diff changeset
    95
ML {*
Christian Urban <urbanc@in.tum.de>
parents: 810
diff changeset
    96
test_funs absF @{context} 
Christian Urban <urbanc@in.tum.de>
parents: 810
diff changeset
    97
     (@{typ "((nat * nat) list) list"}, 
Christian Urban <urbanc@in.tum.de>
parents: 810
diff changeset
    98
      @{typ "((myint) fset) fset"})
Christian Urban <urbanc@in.tum.de>
parents: 810
diff changeset
    99
*}
808
90bde96f5dd1 proper handling of error messages (code copy - maybe this can be avoided)
Christian Urban <urbanc@in.tum.de>
parents: 807
diff changeset
   100
795
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   101
ML {*
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   102
test_funs absF @{context} 
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   103
     (@{typ "(('a * 'a) list) list"}, 
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   104
      @{typ "(('a * 'a) fset) fset"})
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   105
*}
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   106
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   107
ML {*
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   108
test_funs absF @{context} 
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   109
      (@{typ "(nat * nat) list"}, 
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   110
       @{typ "myint fset"})
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   111
*}
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   112
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   113
ML {*
a28f805df355 renamed mk_resp_arg to equiv_relation and exported it in the signature; added tests in AbsRepFun.thy
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   114
test_funs absF @{context} 
796
64f9c76f70c7 corrected wrong [quot_respect] attribute; tuned
Christian Urban <urbanc@in.tum.de>
parents: 795
diff changeset
   115
     (@{typ "('a list) list \<Rightarrow> 'a list"}, 
64f9c76f70c7 corrected wrong [quot_respect] attribute; tuned
Christian Urban <urbanc@in.tum.de>
parents: 795
diff changeset
   116
      @{typ "('a fset) fset \<Rightarrow> 'a fset"})
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   117
*}
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 787
diff changeset
   118
822
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   119
lemma OO_sym_inv:
809
e9e0d1810217 Trying to state composition quotient.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 808
diff changeset
   120
  assumes sr: "symp r"
e9e0d1810217 Trying to state composition quotient.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 808
diff changeset
   121
  and     ss: "symp s"
e9e0d1810217 Trying to state composition quotient.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 808
diff changeset
   122
  shows "(r OO s) x y = (s OO r) y x"
822
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   123
  using sr ss
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   124
  unfolding symp_def
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   125
  apply (metis pred_comp.intros pred_compE ss symp_def)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   126
  done
809
e9e0d1810217 Trying to state composition quotient.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 808
diff changeset
   127
812
18b71981c1f0 Trying the proof
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 810
diff changeset
   128
lemma abs_o_rep:
18b71981c1f0 Trying the proof
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 810
diff changeset
   129
  assumes a: "Quotient r absf repf"
18b71981c1f0 Trying the proof
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 810
diff changeset
   130
  shows "absf o repf = id"
18b71981c1f0 Trying the proof
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 810
diff changeset
   131
  apply(rule ext)
18b71981c1f0 Trying the proof
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 810
diff changeset
   132
  apply(simp add: Quotient_abs_rep[OF a])
822
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   133
  done
812
18b71981c1f0 Trying the proof
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 810
diff changeset
   134
824
cedfe2a71298 Proved concat_empty.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 823
diff changeset
   135
lemma set_in_eq: "(\<forall>e. ((e \<in> A) \<longleftrightarrow> (e \<in> B))) \<equiv> A = B"
822
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   136
  apply (rule eq_reflection)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   137
  apply auto
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   138
  done
816
5edb6facc833 Further in the proof
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 813
diff changeset
   139
822
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   140
lemma map_rel_cong: "b \<approx>1 ba \<Longrightarrow> map f b \<approx>1 map f ba"
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   141
  unfolding erel1_def
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   142
  apply(simp only: set_map set_in_eq)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   143
  done
821
c2ebb7fcf9d0 First generalization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 820
diff changeset
   144
c2ebb7fcf9d0 First generalization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 820
diff changeset
   145
lemma quotient_compose_list_gen_pre:
c2ebb7fcf9d0 First generalization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 820
diff changeset
   146
  assumes a: "equivp r2"
c2ebb7fcf9d0 First generalization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 820
diff changeset
   147
  and b: "Quotient r2 abs2 rep2"
c2ebb7fcf9d0 First generalization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 820
diff changeset
   148
  shows  "(list_rel r2 OO op \<approx>1 OO list_rel r2) r s =
c2ebb7fcf9d0 First generalization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 820
diff changeset
   149
          ((list_rel r2 OO op \<approx>1 OO list_rel r2) r r \<and>
c2ebb7fcf9d0 First generalization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 820
diff changeset
   150
           (list_rel r2 OO op \<approx>1 OO list_rel r2) s s \<and>
c2ebb7fcf9d0 First generalization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 820
diff changeset
   151
           abs_fset (map abs2 r) = abs_fset (map abs2 s))"
822
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   152
  apply rule
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   153
  apply rule
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   154
  apply rule
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   155
  apply (rule list_rel_refl)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   156
  apply (metis equivp_def a)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   157
  apply rule
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   158
  apply (rule equivp_reflp[OF fset_equivp])
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   159
  apply (rule list_rel_refl)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   160
  apply (metis equivp_def a)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   161
  apply(rule)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   162
  apply rule
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   163
  apply (rule list_rel_refl)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   164
  apply (metis equivp_def a)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   165
  apply rule
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   166
  apply (rule equivp_reflp[OF fset_equivp])
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   167
  apply (rule list_rel_refl)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   168
  apply (metis equivp_def a)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   169
  apply (subgoal_tac "map abs2 r \<approx>1 map abs2 s")
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   170
  apply (metis Quotient_rel[OF Quotient_fset])
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   171
  apply (auto)[1]
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   172
  apply (subgoal_tac "map abs2 r = map abs2 b")
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   173
  prefer 2
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   174
  apply (metis Quotient_rel[OF list_quotient[OF b]])
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   175
  apply (subgoal_tac "map abs2 s = map abs2 ba")
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   176
  prefer 2
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   177
  apply (metis Quotient_rel[OF list_quotient[OF b]])
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   178
  apply (simp add: map_rel_cong)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   179
  apply rule
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   180
  apply (rule rep_abs_rsp[of "list_rel r2" "map abs2"])
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   181
  apply (rule list_quotient)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   182
  apply (rule b)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   183
  apply (rule list_rel_refl)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   184
  apply (metis equivp_def a)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   185
  apply rule
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   186
  prefer 2
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   187
  apply (rule rep_abs_rsp_left[of "list_rel r2" "map abs2"])
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   188
  apply (rule list_quotient)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   189
  apply (rule b)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   190
  apply (rule list_rel_refl)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   191
  apply (metis equivp_def a)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   192
  apply (erule conjE)+
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   193
  apply (subgoal_tac "map abs2 r \<approx>1 map abs2 s")
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   194
  apply (rule map_rel_cong)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   195
  apply (assumption)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   196
  apply (metis Quotient_def Quotient_fset equivp_reflp fset_equivp a b)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   197
  done
821
c2ebb7fcf9d0 First generalization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 820
diff changeset
   198
820
162e38c14f24 The working proof of the special case.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 819
diff changeset
   199
lemma quotient_compose_list_gen:
821
c2ebb7fcf9d0 First generalization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 820
diff changeset
   200
  assumes a: "Quotient r2 abs2 rep2"
823
ae3ed7a68e80 Replacing equivp by reflp in the assumptions leads to non-provable subgoals in the gen_pre lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 822
diff changeset
   201
  and     b: "equivp r2" (* reflp is not enough *)
818
9ab49dc166d6 More cleaning and commenting AbsRepTest. Now tests work; just slow.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 817
diff changeset
   202
  shows  "Quotient ((list_rel r2) OO (op \<approx>1) OO (list_rel r2))
811
Christian Urban <urbanc@in.tum.de>
parents: 810
diff changeset
   203
               (abs_fset \<circ> (map abs2)) ((map rep2) \<circ> rep_fset)"
812
18b71981c1f0 Trying the proof
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 810
diff changeset
   204
  unfolding Quotient_def comp_def
822
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   205
  apply (rule)+
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   206
  apply (simp add: abs_o_rep[OF a] id_simps Quotient_abs_rep[OF Quotient_fset])
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   207
  apply (rule)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   208
  apply (rule)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   209
  apply (rule)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   210
  apply (rule list_rel_refl)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   211
  apply (metis b equivp_def)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   212
  apply (rule)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   213
  apply (rule equivp_reflp[OF fset_equivp])
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   214
  apply (rule list_rel_refl)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   215
  apply (metis b equivp_def)
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   216
  apply rule
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   217
  apply rule
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   218
  apply(rule quotient_compose_list_gen_pre[OF b a])
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   219
  done
809
e9e0d1810217 Trying to state composition quotient.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 808
diff changeset
   220
821
c2ebb7fcf9d0 First generalization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 820
diff changeset
   221
(* This is the general statement but the types of abs2 and rep2
c2ebb7fcf9d0 First generalization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 820
diff changeset
   222
   are wrong as can be seen in following exanples *)
818
9ab49dc166d6 More cleaning and commenting AbsRepTest. Now tests work; just slow.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 817
diff changeset
   223
lemma quotient_compose_general:
821
c2ebb7fcf9d0 First generalization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 820
diff changeset
   224
  assumes a2: "Quotient r1 abs1 rep1"
810
f73e2f5f17b2 Struggling with composition
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 809
diff changeset
   225
  and         "Quotient r2 abs2 rep2"
818
9ab49dc166d6 More cleaning and commenting AbsRepTest. Now tests work; just slow.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 817
diff changeset
   226
  shows  "Quotient ((list_rel r2) OO r1 OO (list_rel r2))
821
c2ebb7fcf9d0 First generalization.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 820
diff changeset
   227
               (abs1 \<circ> (map abs2)) ((map rep2) \<circ> rep1)"
810
f73e2f5f17b2 Struggling with composition
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 809
diff changeset
   228
sorry
f73e2f5f17b2 Struggling with composition
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 809
diff changeset
   229
822
5527430f9b6f some cleaning.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 821
diff changeset
   230
thm quotient_compose_list_gen[OF Quotient_fset fset_equivp]
818
9ab49dc166d6 More cleaning and commenting AbsRepTest. Now tests work; just slow.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 817
diff changeset
   231
thm quotient_compose_general[OF Quotient_fset]
817
11a23fe266c4 cleaning in AbsRepTest.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 816
diff changeset
   232
(* Doesn't work: *)
818
9ab49dc166d6 More cleaning and commenting AbsRepTest. Now tests work; just slow.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 817
diff changeset
   233
(* thm quotient_compose_general[OF Quotient_fset Quotient_fset] *)
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
   234
809
e9e0d1810217 Trying to state composition quotient.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 808
diff changeset
   235
end