| author | Christian Urban <urbanc@in.tum.de> | 
| Thu, 31 May 2012 12:01:01 +0100 | |
| changeset 3181 | ca162f0a7957 | 
| parent 1935 | 266abc3ee228 | 
| permissions | -rw-r--r-- | 
| 
714
 
37f7dc85b61b
Added FSet3 with a formalisation of finite sets based on Michael's one.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
1  | 
theory FSet3  | 
| 
1914
 
c50601bc617e
Remove the part already in FSet and leave the experiments
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
1873 
diff
changeset
 | 
2  | 
imports "../../../Nominal/FSet"  | 
| 
714
 
37f7dc85b61b
Added FSet3 with a formalisation of finite sets based on Michael's one.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
3  | 
begin  | 
| 
 
37f7dc85b61b
Added FSet3 with a formalisation of finite sets based on Michael's one.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
4  | 
|
| 
768
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
5  | 
(* TBD *)  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
6  | 
|
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
7  | 
text {* syntax for fset comprehensions (adapted from lists) *}
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
8  | 
|
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
9  | 
nonterminals fsc_qual fsc_quals  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
10  | 
|
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
11  | 
syntax  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
12  | 
"_fsetcompr" :: "'a \<Rightarrow> fsc_qual \<Rightarrow> fsc_quals \<Rightarrow> 'a fset"  ("{|_ . __")
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
13  | 
"_fsc_gen" :: "'a \<Rightarrow> 'a fset \<Rightarrow> fsc_qual" ("_ <- _")
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
14  | 
"_fsc_test" :: "bool \<Rightarrow> fsc_qual" ("_")
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
15  | 
"_fsc_end" :: "fsc_quals" ("|}")
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
16  | 
"_fsc_quals" :: "fsc_qual \<Rightarrow> fsc_quals \<Rightarrow> fsc_quals" (", __")
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
17  | 
"_fsc_abs" :: "'a => 'b fset => 'b fset"  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
18  | 
|
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
19  | 
syntax (xsymbols)  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
20  | 
"_fsc_gen" :: "'a \<Rightarrow> 'a fset \<Rightarrow> fsc_qual" ("_ \<leftarrow> _")
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
21  | 
syntax (HTML output)  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
22  | 
"_fsc_gen" :: "'a \<Rightarrow> 'a fset \<Rightarrow> fsc_qual" ("_ \<leftarrow> _")
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
23  | 
|
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
24  | 
parse_translation (advanced) {*
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
25  | 
let  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
26  | 
  val femptyC = Syntax.const @{const_name fempty};
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
27  | 
  val finsertC = Syntax.const @{const_name finsert};
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
28  | 
  val fmapC = Syntax.const @{const_name fmap};
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
29  | 
  val fconcatC = Syntax.const @{const_name fconcat};
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
30  | 
  val IfC = Syntax.const @{const_name If};
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
31  | 
fun fsingl x = finsertC $ x $ femptyC;  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
32  | 
|
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
33  | 
fun pat_tr ctxt p e opti = (* %x. case x of p => e | _ => [] *)  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
34  | 
let  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
35  | 
val x = Free (Name.variant (fold Term.add_free_names [p, e] []) "x", dummyT);  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
36  | 
val e = if opti then fsingl e else e;  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
37  | 
val case1 = Syntax.const "_case1" $ p $ e;  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
38  | 
val case2 = Syntax.const "_case1" $ Syntax.const Term.dummy_patternN  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
39  | 
$ femptyC;  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
40  | 
val cs = Syntax.const "_case2" $ case1 $ case2  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
41  | 
val ft = Datatype_Case.case_tr false Datatype.info_of_constr  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
42  | 
ctxt [x, cs]  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
43  | 
in lambda x ft end;  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
44  | 
|
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
45  | 
fun abs_tr ctxt (p as Free(s,T)) e opti =  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
46  | 
let val thy = ProofContext.theory_of ctxt;  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
47  | 
val s' = Sign.intern_const thy s  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
48  | 
in if Sign.declared_const thy s'  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
49  | 
then (pat_tr ctxt p e opti, false)  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
50  | 
else (lambda p e, true)  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
51  | 
end  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
52  | 
| abs_tr ctxt p e opti = (pat_tr ctxt p e opti, false);  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
53  | 
|
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
54  | 
  fun fsc_tr ctxt [e, Const("_fsc_test",_) $ b, qs] =
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
55  | 
let  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
56  | 
val res = case qs of  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
57  | 
                      Const("_fsc_end",_) => fsingl e
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
58  | 
                    | Const("_fsc_quals",_)$ q $ qs => fsc_tr ctxt [e, q, qs];
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
59  | 
in  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
60  | 
IfC $ b $ res $ femptyC  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
61  | 
end  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
62  | 
|
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
63  | 
    | fsc_tr ctxt [e, Const("_fsc_gen",_) $ p $ es, Const("_fsc_end",_)] =
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
64  | 
(case abs_tr ctxt p e true of  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
65  | 
(f,true) => fmapC $ f $ es  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
66  | 
| (f, false) => fconcatC $ (fmapC $ f $ es))  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
67  | 
|
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
68  | 
    | fsc_tr ctxt [e, Const("_fsc_gen",_) $ p $ es, Const("_fsc_quals",_) $ q $ qs] =
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
69  | 
let  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
70  | 
val e' = fsc_tr ctxt [e, q, qs];  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
71  | 
in  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
72  | 
fconcatC $ (fmapC $ (fst (abs_tr ctxt p e' false)) $ es)  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
73  | 
end  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
74  | 
|
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
75  | 
in [("_fsetcompr", fsc_tr)] end
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
76  | 
*}  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
77  | 
|
| 
1260
 
9df6144e281b
moved Quot package to Attic (still compiles there with "isabelle make")
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1144 
diff
changeset
 | 
78  | 
|
| 
 
9df6144e281b
moved Quot package to Attic (still compiles there with "isabelle make")
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1144 
diff
changeset
 | 
79  | 
(* NEEDS FIXING *)  | 
| 
768
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
80  | 
(* examles *)  | 
| 
1260
 
9df6144e281b
moved Quot package to Attic (still compiles there with "isabelle make")
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1144 
diff
changeset
 | 
81  | 
(*  | 
| 
768
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
82  | 
term "{|(x,y,z). b|}"
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
83  | 
term "{|x. x \<leftarrow> xs|}"
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
84  | 
term "{|(x,y,z). x\<leftarrow>xs|}"
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
85  | 
term "{|e x y. x\<leftarrow>xs, y\<leftarrow>ys|}"
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
86  | 
term "{|(x,y,z). x<a, x>b|}"
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
87  | 
term "{|(x,y,z). x\<leftarrow>xs, x>b|}"
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
88  | 
term "{|(x,y,z). x<a, x\<leftarrow>xs|}"
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
89  | 
term "{|(x,y). Cons True x \<leftarrow> xs|}"
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
90  | 
term "{|(x,y,z). Cons x [] \<leftarrow> xs|}"
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
91  | 
term "{|(x,y,z). x<a, x>b, x=d|}"
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
92  | 
term "{|(x,y,z). x<a, x>b, y\<leftarrow>ys|}"
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
93  | 
term "{|(x,y,z). x<a, x\<leftarrow>xs,y>b|}"
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
94  | 
term "{|(x,y,z). x<a, x\<leftarrow>xs, y\<leftarrow>ys|}"
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
95  | 
term "{|(x,y,z). x\<leftarrow>xs, x>b, y<a|}"
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
96  | 
term "{|(x,y,z). x\<leftarrow>xs, x>b, y\<leftarrow>ys|}"
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
97  | 
term "{|(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,y>x|}"
 | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
98  | 
term "{|(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,z\<leftarrow>zs|}"
 | 
| 
1260
 
9df6144e281b
moved Quot package to Attic (still compiles there with "isabelle make")
 
Christian Urban <urbanc@in.tum.de> 
parents: 
1144 
diff
changeset
 | 
99  | 
*)  | 
| 
768
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
100  | 
|
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
101  | 
(* BELOW CONSTRUCTION SITE *)  | 
| 
 
e9e205b904e2
get_fun needed change to cope with "('a fset) fset" types...this needs composition (op o); now id_simps contains also id_o and o_id, and map_id is also added in QuotList.thy; regularize and cleaning needed to be hacked (indicated by "HACK")...THIS NEEDS ATTENTION!!!; except two lemmas in IntEx, all examples go through; added considerable material to FSet3; tuned FIXME-TODO
 
Christian Urban <urbanc@in.tum.de> 
parents: 
767 
diff
changeset
 | 
102  | 
|
| 
714
 
37f7dc85b61b
Added FSet3 with a formalisation of finite sets based on Michael's one.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
103  | 
|
| 
 
37f7dc85b61b
Added FSet3 with a formalisation of finite sets based on Michael's one.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
104  | 
end  |