diff -r 13f985a93dbc -r 84a356e3d38b QuotTest.thy --- a/QuotTest.thy Wed Oct 28 19:36:52 2009 +0100 +++ b/QuotTest.thy Wed Oct 28 19:46:15 2009 +0100 @@ -81,21 +81,21 @@ *} ML {* - get_fun repF @{typ t} @{typ qt} @{context} @{typ "((((qt \ qt) \ qt) \ qt) list) * nat"} + get_fun repF [(@{typ qt}, @{typ qt})] @{context} @{typ "((((qt \ qt) \ qt) \ qt) list) * nat"} |> fst |> Syntax.string_of_term @{context} |> writeln *} ML {* - get_fun absF @{typ t} @{typ qt} @{context} @{typ "qt * nat"} + get_fun absF [(@{typ qt}, @{typ t})] @{context} @{typ "qt * nat"} |> fst |> Syntax.string_of_term @{context} |> writeln *} ML {* - get_fun absF @{typ t} @{typ qt} @{context} @{typ "(qt \ qt) \ qt"} + get_fun absF [(@{typ qt}, @{typ t})] @{context} @{typ "(qt \ qt) \ qt"} |> fst |> Syntax.pretty_term @{context} |> Pretty.string_of @@ -111,11 +111,20 @@ *} *) -local_setup {* - make_const_def @{binding VR} @{term "vr"} NoSyn @{typ "t"} @{typ "qt"} #> snd #> - make_const_def @{binding AP} @{term "ap"} NoSyn @{typ "t"} @{typ "qt"} #> snd #> - make_const_def @{binding LM} @{term "lm"} NoSyn @{typ "t"} @{typ "qt"} #> snd -*} +quotient_def (for qt) + VR ::"string \ qt" +where + "VR \ vr" + +quotient_def (for qt) + AP ::"qt list \ qt" +where + "AP \ ap" + +quotient_def (for qt) + LM ::"string \ qt \ qt" +where + "LM \ lm" term vr term ap @@ -141,11 +150,21 @@ print_quotients print_theorems -local_setup {* - make_const_def @{binding VR'} @{term "vr'"} NoSyn @{typ "'a t'"} @{typ "'a qt'"} #> snd #> - make_const_def @{binding AP'} @{term "ap'"} NoSyn @{typ "'a t'"} @{typ "'a qt'"} #> snd #> - make_const_def @{binding LM'} @{term "lm'"} NoSyn @{typ "'a t'"} @{typ "'a qt'"} #> snd -*} + +quotient_def (for "'a qt'") + VR' ::"string \ 'a qt" +where + "VR' \ vr'" + +quotient_def (for "'a qt'") + AP' ::"('a qt') * ('a qt') \ 'a qt" +where + "AP' \ ap'" + +quotient_def (for "'a qt'") + LM' ::"'a \ string \ 'a qt'" +where + "LM' \ lm'" term vr' term ap'