let regexp nlet_corec = "nlet" utf8_blank+ "corec"
let regexp ident_decoration = '\'' | '?' | '`'
let regexp ident_cont = ident_letter | xml_digit | '_'
-let regexp ident = ident_letter ident_cont* ident_decoration*
+let regexp ident_start = ident_letter
+let regexp ident = ident_letter ident_cont* ident_decoration*
+let regexp variable_ident = '_' '_' number
+let regexp pident = '_' ident
let regexp tex_token = '\\' ident
let regexp delim_begin = "\\["
let regexp delim_end = "\\]"
-let regexp qkeyword = "'" ident "'"
+let regexp qkeyword = "'" ( ident | pident ) "'"
let regexp implicit = '?'
let regexp placeholder = '%'
let end_cnum = level2_pattern_token_group 0 buffer lexbuf in
return_with_loc (token_name, Buffer.contents buffer) begin_cnum end_cnum
+let handle_keywords lexbuf k name =
+ let s = Ulexing.utf8_lexeme lexbuf in
+ if k s then
+ return lexbuf ("", s)
+ else
+ return lexbuf (name, s)
+;;
+
let rec level2_meta_token =
lexer
| utf8_blank+ -> level2_meta_token lexbuf
| ident ->
- let s = Ulexing.utf8_lexeme lexbuf in
- begin
- if List.mem s level2_meta_keywords then
- return lexbuf ("", s)
- else
- return lexbuf ("IDENT", s)
- end
+ handle_keywords lexbuf (fun x -> List.mem x level2_meta_keywords) "IDENT"
+ | variable_ident -> return lexbuf ("IDENT", Ulexing.utf8_lexeme lexbuf)
+ | pident ->
+ handle_keywords lexbuf (fun x -> List.mem x level2_meta_keywords) "PIDENT"
| "@{" -> read_unparsed_group "UNPARSED_AST" lexbuf
| ast_ident ->
return lexbuf ("UNPARSED_AST",
return lexbuf ("META", String.sub s 1 (String.length s - 1))
| implicit -> return lexbuf ("IMPLICIT", "")
| placeholder -> return lexbuf ("PLACEHOLDER", "")
- | ident ->
- let lexeme = Ulexing.utf8_lexeme lexbuf in
- if Hashtbl.mem !level2_ast_keywords lexeme then
- return lexbuf ("", lexeme)
- else
- return lexbuf ("IDENT", lexeme)
+ | ident -> handle_keywords lexbuf (Hashtbl.mem !level2_ast_keywords) "IDENT"
+ | variable_ident -> return lexbuf ("IDENT", Ulexing.utf8_lexeme lexbuf)
+ | pident -> handle_keywords lexbuf (Hashtbl.mem !level2_ast_keywords) "PIDENT"
| number -> return lexbuf ("NUMBER", Ulexing.utf8_lexeme lexbuf)
| tex_token -> return lexbuf (expand_macro lexbuf)
| uri -> return lexbuf ("URI", Ulexing.utf8_lexeme lexbuf)
lexer
| utf8_blank+ -> ligatures_token level1_pattern_token lexbuf
| number -> return lexbuf ("NUMBER", Ulexing.utf8_lexeme lexbuf)
- | ident ->
- let s = Ulexing.utf8_lexeme lexbuf in
- begin
- if List.mem s level1_keywords then
- return lexbuf ("", s)
- else
- return lexbuf ("IDENT", s)
- end
+ | ident ->handle_keywords lexbuf (fun x -> List.mem x level1_keywords) "IDENT"
+ | variable_ident -> return lexbuf ("IDENT", Ulexing.utf8_lexeme lexbuf)
+ | pident->handle_keywords lexbuf (fun x->List.mem x level1_keywords) "PIDENT"
| color -> return lexbuf ("COLOR", Ulexing.utf8_lexeme lexbuf)
| percentage ->
return lexbuf ("PERCENTAGE", Ulexing.utf8_lexeme lexbuf)
[ LPAREN; id = single_arg; SYMBOL ":"; typ = term; RPAREN ->
id, Some typ
| arg = single_arg -> arg, None
+ | id = PIDENT -> Ast.Ident (id, None), None
| SYMBOL "_" -> Ast.Ident ("_", None), None
+ | LPAREN; id = PIDENT; SYMBOL ":"; typ = term; RPAREN ->
+ Ast.Ident (id, None), Some typ
| LPAREN; SYMBOL "_"; SYMBOL ":"; typ = term; RPAREN ->
Ast.Ident ("_", None), Some typ
]
];
match_pattern: [
- [ id = IDENT -> Ast.Pattern (id, None, [])
+ [ SYMBOL "_" -> Ast.Wildcard
+ | id = IDENT -> Ast.Pattern (id, None, [])
| LPAREN; id = IDENT; vars = LIST1 possibly_typed_name; RPAREN ->
Ast.Pattern (id, None, vars)
| id = IDENT; vars = LIST1 possibly_typed_name ->
Ast.Pattern (id, None, vars)
- | SYMBOL "_" -> Ast.Wildcard
]
];
binder: [
]
];
binder_vars: [
- [ vars = [
- l = LIST1 single_arg SEP SYMBOL "," -> l
- | SYMBOL "_" -> [Ast.Ident ("_", None)] ];
+ [ vars = [ l =
+ [ l = LIST1 single_arg SEP SYMBOL "," -> l
+ | l = LIST1 [ PIDENT | SYMBOL "_" ] SEP SYMBOL "," ->
+ List.map (fun x -> Ast.Ident(x,None)) l
+ ] -> l ];
typ = OPT [ SYMBOL ":"; t = term -> t ] -> (vars, typ)
]
];
NTactics.rewrite_tac ~dir ~what:(text,prefix_len,what)
~where:(text,prefix_len,where)
;;
+
+let subst_metasenv_and_fix_names s =
+ let u,h,metasenv, subst,o = s.NTacStatus.istatus.NTacStatus.pstatus in
+ let o =
+ NCicUntrusted.map_obj_kind ~skip_body:true
+ (NCicUntrusted.apply_subst subst []) o
+ in
+ { s with NTacStatus.istatus =
+ { s.NTacStatus.istatus with NTacStatus.pstatus =
+ u,h,NCicUntrusted.apply_subst_metasenv subst metasenv,subst,o}}
+;;
let rec eval_command = {ec_go = fun ~disambiguate_command opts status
(text,prefix_len,cmd) ->
let obj =
prerr_endline "CSC: here we should fix the height!!!";
(uri,height,[],[],
- NCicUntrusted.map_obj_kind (NCicUntrusted.apply_subst subst)
+ NCicUntrusted.map_obj_kind
+ (NCicUntrusted.apply_subst subst [])
obj) in
NCicTypeChecker.typecheck_obj obj;
NCicLibrary.add_obj uri obj;
{ status with
GrafiteTypes.ng_status =
GrafiteTypes.ProofMode
- { NTacStatus.gstatus = ninitial_stack;
- istatus = { NTacStatus.pstatus = obj; lstatus = lexicon_status}}
- }
+ (subst_metasenv_and_fix_names
+ { NTacStatus.gstatus = ninitial_stack;
+ istatus = { NTacStatus.pstatus = obj; lstatus = lexicon_status}})
+ }
in
(match nmenv with
[] ->
| GrafiteTypes.CommandMode _ -> assert false
| GrafiteTypes.ProofMode nstatus ->
let nstatus = eval_ng_tac (text,prefix_len,tac) nstatus in
+ let nstatus = subst_metasenv_and_fix_names nstatus in
let nstatus = eval_ng_punct (text,prefix_len,punct) nstatus in
NTacStatus.pp_tac_status nstatus;
{ status with GrafiteTypes.ng_status= GrafiteTypes.ProofMode nstatus },
constructor: [ [ name = IDENT; SYMBOL ":"; typ = term -> (name, typ) ] ];
tactic_term: [ [ t = term LEVEL "90" -> t ] ];
new_name: [
- [ id = IDENT -> Some id
- | SYMBOL "_" -> None ]
+ [ SYMBOL "_" -> None
+ | id = IDENT -> Some id ]
];
ident_list0: [ [ LPAREN; idents = LIST0 new_name; RPAREN -> idents ] ];
tactic_term_list1: [
];
level3_term: [
[ u = URI -> N.UriPattern (UriManager.uri_of_string u)
+ | IMPLICIT -> N.ImplicitPattern
| id = IDENT -> N.VarPattern id
- | SYMBOL "_" -> N.ImplicitPattern
| LPAREN; terms = LIST1 SELF; RPAREN ->
(match terms with
| [] -> assert false
let skip s = true in
let inline s = List.mem s [ "int" ] in
- let rec visit_entry e todo is_son =
+ let rec visit_entry e ?level todo is_son =
let { ename = ename; edesc = desc } = e in
if inline ename then
visit_desc desc todo is_son
else
begin
- Format.fprintf fmt "%s " ename;
- if skip ename then
- todo
- else
- todo @ [e]
+ (match level with
+ | None -> Format.fprintf fmt "%s " ename;
+ | Some _ -> Format.fprintf fmt "%s " ename;);
+ if skip ename then
+ todo
+ else
+ todo @ [e]
end
and visit_desc d todo is_son =
todo)
todo sl
| Snterm entry -> visit_entry entry todo is_son
- | Snterml (entry,_) -> visit_entry entry todo is_son
+ | Snterml (entry,level) -> visit_entry entry ~level todo is_son
| Slist0 symbol ->
Format.fprintf fmt "{@[<hov2> ";
let todo = visit_symbol symbol todo is_son in
| NCic.Implicit `Hole -> idref (Ast.UserInput)
| NCic.Implicit _ -> idref (Ast.Implicit)
| NCic.Prod (n,s,t) ->
+ let n = if n.[0] = '_' then "_" else n in
let binder_kind = `Forall in
idref (Ast.Binder (binder_kind, (Ast.Ident (n,None), Some (k ctx s)),
k ((n,NCic.Decl s)::ctx) t))
| _ -> NCic.Appl (he::args)
;;
-let map_obj_kind f =
+let map_obj_kind ?(skip_body=false) f =
+ let do_bo f x = if skip_body then x else f x in
function
NCic.Constant (relev,name,bo,ty,attrs) ->
- NCic.Constant (relev,name,HExtlib.map_option f bo, f ty,attrs)
+ NCic.Constant (relev,name,do_bo (HExtlib.map_option f) bo, f ty,attrs)
| NCic.Fixpoint (ind,fl,attrs) ->
let fl =
List.map
- (function (relevance,name,recno,ty,bo) -> relevance,name,recno,f ty,f bo)
+ (function (relevance,name,recno,ty,bo) ->
+ relevance,name,recno,f ty,do_bo f bo)
fl
in
NCic.Fixpoint (ind,fl,attrs)
NCic.Inductive (is_ind,lno,itl,attrs)
;;
-let apply_subst subst t =
+exception Occurr;;
+
+let clean_or_fix_dependent_abstrations ctx t =
+ let occurrs_1 t =
+ let rec aux n _ = function
+ | NCic.Meta _ -> ()
+ | NCic.Rel i when i = n + 1 -> raise Occurr
+ | t -> NCicUtils.fold (fun _ k -> k + 1) n aux () t
+ in
+ try aux 0 () t; false
+ with Occurr -> true
+ in
+ let fresh ctx name =
+ if not (List.mem name ctx) then name
+ else
+ let rec aux i =
+ let attempt = name ^ string_of_int i in
+ if List.mem attempt ctx then aux (i+1)
+ else attempt
+ in
+ aux 0
+ in
+ let rec aux ctx = function
+ | NCic.Meta _ as t -> t
+ | NCic.Prod (name,s,t) when name.[0] = '#' && occurrs_1 t ->
+ let name = fresh ctx (String.sub name 1 (String.length name-1)) in
+ NCic.Prod (name,aux ctx s, aux (name::ctx) t)
+ | NCic.Prod (name,s,t) when name.[0] = '#' && not (occurrs_1 t) ->
+ NCic.Prod ("_",aux ctx s,aux ("_"::ctx) t)
+ | NCic.Prod ("_",s,t) -> NCic.Prod("_",aux ctx s,aux ("_"::ctx) t)
+ | NCic.Prod (name,s,t) when name.[0] <> '_' && not (occurrs_1 t) ->
+ let name = fresh ctx ("_"^name) in
+ NCic.Prod (name, aux ctx s, aux (name::ctx) t)
+ | NCic.Prod (name,s,t) when List.mem name ctx ->
+ let name = fresh ctx name in
+ NCic.Prod (name, aux ctx s, aux (name::ctx) t)
+ | NCic.Lambda (name,s,t) when List.mem name ctx ->
+ let name = fresh ctx name in
+ NCic.Lambda (name, aux ctx s, aux (name::ctx) t)
+ | t -> NCicUtils.map (fun (e,_) ctx -> e::ctx) ctx aux t
+ in
+ aux (List.map fst ctx) t
+;;
+
+let apply_subst subst context t =
let rec apply_subst subst () =
function
NCic.Meta (i,lc) ->
apply_subst subst () (NCicSubstitution.lift n t)) l))))
| t -> NCicUtils.map (fun _ () -> ()) () (apply_subst subst) t
in
- apply_subst subst () t
+ clean_or_fix_dependent_abstrations context (apply_subst subst () t)
;;
+let apply_subst_context subst context =
+ let rec aux c = function
+ | [] -> []
+ | (name,NCic.Decl t as e) :: tl ->
+ (name, NCic.Decl (apply_subst subst c t)) :: aux (e::c) tl
+ | (name,NCic.Def (t1,t2) as e) :: tl ->
+ (name, NCic.Def (apply_subst subst c t1,apply_subst subst c t2)) ::
+ aux (e::c) tl
+ in
+ List.rev (aux [] (List.rev context))
+;;
+
+let rec apply_subst_metasenv subst = function
+ | [] -> []
+ | (i,_) :: _ when List.mem_assoc i subst -> assert false
+ | (i,(name,ctx,ty)) :: tl ->
+ (i,(name,apply_subst_context subst ctx,apply_subst subst ctx ty)) ::
+ apply_subst_metasenv subst tl
+;;
(NCic.hypothesis -> 'k -> 'k) -> 'k ->
('k -> 'a -> NCic.term -> 'a * NCic.term) -> 'a -> NCic.term -> 'a * NCic.term
-val map_obj_kind: (NCic.term -> NCic.term) -> NCic.obj_kind -> NCic.obj_kind
+val map_obj_kind:
+ ?skip_body:bool -> (NCic.term -> NCic.term) -> NCic.obj_kind -> NCic.obj_kind
val metas_of_term : NCic.substitution -> NCic.context -> NCic.term -> int list
val mk_appl : NCic.term -> NCic.term list -> NCic.term
-val apply_subst : NCic.substitution -> NCic.term -> NCic.term
+(* the context is needed only to honour Barendregt's naming convention *)
+val apply_subst : NCic.substitution -> NCic.context -> NCic.term -> NCic.term
+val apply_subst_metasenv : NCic.substitution -> NCic.metasenv -> NCic.metasenv
(NCicPp.ppterm ~subst ~metasenv ~context y)
(NCicPp.ppterm ~subst ~metasenv ~context x))))
+and guess_name subst ctx ty =
+ let aux initial = "#" ^ String.make 1 initial in
+ match ty with
+ | C.Const (NReference.Ref (u,_))
+ | C.Appl (C.Const (NReference.Ref (u,_)) :: _) ->
+ aux (String.sub (NUri.name_of_uri u) 0 1).[0]
+ | C.Prod _ -> aux 'f'
+ | C.Meta (n,lc) ->
+ (try
+ let _,_,t,_ = NCicUtils.lookup_subst n subst in
+ guess_name subst ctx (NCicSubstitution.subst_meta lc t)
+ with NCicUtils.Subst_not_found _ -> aux 'M')
+ | _ -> aux 'H'
+
and eat_prods hdb
~localise ~look_for_coercion metasenv subst context orig_he he ty_he args
=
let metasenv, subst, arg, ty_arg =
typeof hdb ~look_for_coercion ~localise
metasenv subst context arg None in
+ let name = guess_name subst context ty_arg in
let metasenv, _, meta, _ =
NCicMetaSubst.mk_meta metasenv
- (("_",C.Decl ty_arg) :: context) `Type
+ ((name,C.Decl ty_arg) :: context) `Type
in
- let flex_prod = C.Prod ("_", ty_arg, meta) in
+ let flex_prod = C.Prod (name, ty_arg, meta) in
(* next line grants that ty_args is a type *)
let metasenv,subst, flex_prod, _ =
typeof hdb ~look_for_coercion ~localise metasenv subst
let apply_subst status ctx t =
let status, (name,_,t) = relocate status ctx t in
let _,_,_,subst,_ = status.pstatus in
- status, (name, ctx, NCicUntrusted.apply_subst subst t)
+ status, (name, ctx, NCicUntrusted.apply_subst subst ctx t)
;;
| Some eq_uri ->
try
let bag, current = Equality.equality_of_term bag term ty newmetas in
- let bag, current = Equality.fix_metas bag current in
- match add_to_active_aux bag active passive env eq_uri current with
- | _,a,p,b -> a,p,b
+ let w,_,_,_,_ = Equality.open_equality current in
+ if w > 100 then
+ (HLog.debug
+ ("skipping giant " ^ CicPp.ppterm term ^ " of weight " ^
+ string_of_int w); active, passive, bag)
+ else
+ let bag, current = Equality.fix_metas bag current in
+ match add_to_active_aux bag active passive env eq_uri current with
+ | _,a,p,b -> a,p,b
with
| Equality.TermIsNotAnEquality -> active, passive, bag
;;
applyTransformation.cmx matitaAutoGui.cmi
matitaclean.cmo: matitaMisc.cmi matitaInit.cmi matitaclean.cmi
matitaclean.cmx: matitaMisc.cmx matitaInit.cmx matitaclean.cmi
-matitacLib.cmo: matitaExcPp.cmi matitaEngine.cmi buildTimeConf.cmo \
- applyTransformation.cmi matitacLib.cmi
-matitacLib.cmx: matitaExcPp.cmx matitaEngine.cmx buildTimeConf.cmx \
- applyTransformation.cmx matitacLib.cmi
+matitacLib.cmo: matitaMisc.cmi matitaExcPp.cmi matitaEngine.cmi \
+ buildTimeConf.cmo applyTransformation.cmi matitacLib.cmi
+matitacLib.cmx: matitaMisc.cmx matitaExcPp.cmx matitaEngine.cmx \
+ buildTimeConf.cmx applyTransformation.cmx matitacLib.cmi
matitac.cmo: matitadep.cmi matitaclean.cmi matitacLib.cmi matitaWiki.cmo \
matitaMisc.cmi matitaInit.cmi
matitac.cmx: matitadep.cmx matitaclean.cmx matitacLib.cmx matitaWiki.cmx \
cs_ap : relation cs_crr;
cs_proof : is_CSetoid cs_crr cs_eq cs_ap}.
-interpretation "setoid equality" 'eq x y = (cs_eq _ x y).
+interpretation "setoid equality" 'eq x y = (cs_eq ? x y).
-interpretation "setoid apart" 'neq x y = (cs_ap _ x y).
+interpretation "setoid apart" 'neq x y = (cs_ap ? x y).
(* visto che sia "ap" che "eq" vanno in Prop e data la "tight-apartness",
"cs_neq" e "ap" non sono la stessa cosa? *)
interpretation "Coq 7.3.1 logical and" 'and x y = (cic:/matita/LAMBDA-TYPES/Legacy-1/coq/defs/land.ind#xpointer(1/1) x y).
interpretation "Coq 7.3.1 logical or" 'or x y = (cic:/matita/LAMBDA-TYPES/Legacy-1/coq/defs/or.ind#xpointer(1/1) x y).
interpretation "Coq 7.3.1 logical not" 'not x = (cic:/matita/LAMBDA-TYPES/Legacy-1/coq/defs/not.con x).
-interpretation "Coq 7.3.1 exists" 'exists \eta.x = (cic:/matita/LAMBDA-TYPES/Legacy-1/coq/defs/ex.ind#xpointer(1/1) _ x).
+interpretation "Coq 7.3.1 exists" 'exists \eta.x = (cic:/matita/LAMBDA-TYPES/Legacy-1/coq/defs/ex.ind#xpointer(1/1) ? x).
interpretation "Coq 7.3.1 natural 'less or equal to'" 'leq x y = (cic:/matita/LAMBDA-TYPES/Legacy-1/coq/defs/le.ind#xpointer(1/1) x y).
interpretation "Coq 7.3.1 natural 'less than'" 'lt x y = (cic:/matita/LAMBDA-TYPES/Legacy-1/coq/defs/lt.con x y).
interpretation "Coq 7.3.1 leibnitz's equality" 'eq t x y = (cic:/matita/LAMBDA-TYPES/Legacy-1/coq/defs/eq.ind#xpointer(1/1) t x y).
notation < "hvbox(x \nbsp ∈ \nbsp l)"
non associative with precedence 30 for @{ 'inlist $x $l }.
*)
-interpretation "item in list" 'mem x l = (in_list _ x l).
+interpretation "item in list" 'mem x l = (in_list ? x l).
lemma max_case : \forall m,n.(max m n) = match (leb m n) with
[ true \Rightarrow n
(*CSC: the URI must disappear: there is a bug now *)
interpretation "nil" 'nil =
- (cic:/matita/RELATIONAL/datatypes/List/List.ind#xpointer(1/1) _).
+ (cic:/matita/RELATIONAL/datatypes/List/List.ind#xpointer(1/1) ?).
notation "hvbox([])"
non associative with precedence 95
(*CSC: the URI must disappear: there is a bug now *)
interpretation "right cons" 'rcons x y =
- (cic:/matita/RELATIONAL/datatypes/List/List.ind#xpointer(1/2) _ x y).
+ (cic:/matita/RELATIONAL/datatypes/List/List.ind#xpointer(1/2) ? x y).
notation "hvbox([a break @ b])"
non associative with precedence 95
notation "\ll term 19 a, break term 19 b \gg"
with precedence 90 for @{'dependent_pair (λx:?.? x) $a $b}.
-interpretation "dependent pair" 'dependent_pair \eta.c a b = (sigma_intro _ c a b).
+interpretation "dependent pair" 'dependent_pair \eta.c a b = (sigma_intro ? c a b).
-interpretation "sigma" 'Sigma \eta.x = (sigma _ x).
+interpretation "sigma" 'Sigma \eta.x = (sigma ? x).
definition sigmaFst ≝
λT:Type.λf:T → Type.λs:sigma T f.match s with [ sigma_intro x _ ⇒ x ].
right associative with precedence 47
for @{'ne_append $l1 $l2 }.
-interpretation "ne_nil" 'ne_nil hd = (ne_nil _ hd).
-interpretation "ne_cons" 'ne_cons hd tl = (ne_cons _ hd tl).
-interpretation "ne_append" 'ne_append l1 l2 = (ne_append _ l1 l2).
+interpretation "ne_nil" 'ne_nil hd = (ne_nil ? hd).
+interpretation "ne_cons" 'ne_cons hd tl = (ne_cons ? hd tl).
+interpretation "ne_append" 'ne_append l1 l2 = (ne_append ? l1 l2).
(* ************ *)
(* List Utility *)
rewrite > mod_plus in ⊢ (? ? % ?);
rewrite > mod_plus in ⊢ (? ? ? %);
apply eq_mod_to_eq_plus_mod;
- rewrite < mod_mod in ⊢ (? ? ? %); [ | autobatch];
- rewrite < mod_mod in ⊢ (? ? % ?); [ | autobatch];
+ rewrite < mod_mod in ⊢ (? ? ? %); [ | autobatch by divides_n_n];
+ rewrite < mod_mod in ⊢ (? ? % ?); [ | autobatch by divides_n_n];
rewrite < (eq_mod_times_times_mod ? ? 16 256) in ⊢ (? ? (? % ?) ?); [2: reflexivity | ];
rewrite < mod_mod in ⊢ (? ? % ?);
[ reflexivity
- | autobatch
+ | autobatch by divides_n_n
].
qed.
rewrite > mod_plus;
letin K ≝ (eq_nat_of_byte8_mod a); clearbody K;
letin K' ≝ (eq_mod_times_n_m_m_O (nat_of_bool b) 256 ?); clearbody K';
- [ autobatch | ];
- autobatch paramodulation.
+ [ unfold;apply le_S_S;apply le_O_n| ];
+ rewrite > K' in ⊢ (? ? ? (? (? ? %) ?));
+ rewrite < K in ⊢ (? ? ? (? (? % ?) ?));
+ rewrite < plus_n_O;
+ apply K;
qed.
lemma eq_eqb8_x0_x0_byte8_of_nat_S_false:
| unfold eq_b8;
change in ⊢ (? ? (? % ?) ?) with (eq_ex x0 (exadecim_of_nat (S b/16)));
letin K ≝ (leq_m_n_to_eq_div_n_m_S (S b) 16 ? ?);
- [ autobatch
+ [ unfold; autobatch by le_S_S,le_O_n;
| unfold in H1;
apply le_S_S;
assumption
unfold lt in H;
rewrite < H2;
clear H2; clear a; clear H1; clear Hcut;
- apply (le_times_to_le 16) [ autobatch | ] ;
- rewrite > (div_mod (S b) 16) in H;[2:autobatch|]
- rewrite > (div_mod 255 16) in H:(? ? %);[2:autobatch|]
+ apply (le_times_to_le 16) [ autobatch by le_S_S,le_O_n;| ] ;
+ rewrite > (div_mod (S b) 16) in H;[2:unfold; autobatch by le_S_S,le_O_n;|]
+ rewrite > (div_mod 255 16) in H:(? ? %);[2:unfold; autobatch by le_S_S,le_O_n;|]
lapply (le_to_le_plus_to_le ? ? ? ? ? H);
[apply lt_S_to_le;
- apply lt_mod_m_m;autobatch
+ apply lt_mod_m_m; unfold; apply le_S_S; apply le_O_n;
|rewrite > sym_times;
rewrite > sym_times in ⊢ (? ? %);
normalize in ⊢ (? ? %);apply Hletin;
rewrite > nat_of_exadecim_exadecim_of_nat in Hletin1;
elim (eq_mod_O_to_exists ? ? Hletin1); clear Hletin1;
rewrite > H1;
- rewrite > lt_O_to_div_times; [2: autobatch | ]
+ rewrite > lt_O_to_div_times; [2: unfold; apply le_S_S; apply le_O_n; | ]
lapply (eq_f ? ? (λx.x/16) ? ? H1);
- rewrite > lt_O_to_div_times in Hletin; [2: autobatch | ]
+ rewrite > lt_O_to_div_times in Hletin; [2: unfold; apply le_S_S; apply le_O_n; | ]
lapply (eq_f ? ? (λx.x \mod 16) ? ? H1);
rewrite > eq_mod_times_n_m_m_O in Hletin1;
elim daemon
rewrite > eq_nat_of_byte8_n_nat_of_byte8_mod_n_256 in ⊢ (? ? ? %);
rewrite > mod_plus in ⊢ (? ? (? %) ?);
rewrite > mod_plus in ⊢ (? ? ? (? %));
- rewrite < mod_mod in ⊢ (? ? (? (? (? % ?) ?)) ?); [2: autobatch | ];
+ rewrite < mod_mod in ⊢ (? ? (? (? (? % ?) ?)) ?); [2: apply divides_n_n; | ];
rewrite > sym_plus in ⊢ (? ? (? (? % ?)) ?);
reflexivity.
qed.
intro;
elim e;
simplify;
- autobatch depth=17.
+ unfold;autobatch depth=17 size=20.
qed.
lemma exadecim_of_nat_mod:
rewrite < mod_mod;
[ apply H;
unfold lt;
- autobatch.
+ autobatch depth=20 size=20.
| autobatch
]
qed.
lemma le_to_lt: ∀n,m. n ≤ m → n < S m.
intros;
- autobatch.
+ unfold;autobatch.
qed.
alias num (instance 0) = "natural number".
]
| elim H1; autobatch
]
- | autobatch
+ | exists;[apply (pred m);]autobatch
].
qed.
| FamilyHCS08 : HCS08_mcu_model → any_mcu_model
| FamilyRS08 : RS08_mcu_model → any_mcu_model.
-coercion FamilyHC05.
-coercion FamilyHC08.
-coercion FamilyHCS08.
-coercion FamilyRS08.
+definition c_FamilyHC05:= FamilyHC05.
+definition c_FamilyHC08:= FamilyHC08.
+definition c_FamilyHCS08:=FamilyHCS08.
+definition c_FamilyRS08:= FamilyRS08.
+
+coercion c_FamilyHC05.
+coercion c_FamilyHC08.
+coercion c_FamilyHCS08.
+coercion c_FamilyRS08.
(*
condizioni errore interne alla MCU
inductive any_opcode (m:mcu_type) : Type ≝
anyOP : opcode → any_opcode m.
-coercion anyOP.
+definition c_anyOP ≝ anyOP.
+
+coercion c_anyOP.
(* raggruppamento di byte e word in un tipo unico *)
inductive byte8_or_word16 : Type ≝
Byte: byte8 → byte8_or_word16
| Word: word16 → byte8_or_word16.
-coercion Byte.
-coercion Word.
+definition c_Byte ≝ Byte.
+definition c_Word ≝ Word.
+
+coercion c_Byte.
+coercion c_Word.
(* opcode → naturali, per usare eqb *)
definition magic_of_opcode ≝
inductive t_byte8 (m:mcu_type) : Type ≝
TByte : byte8 → t_byte8 m.
-coercion TByte.
+definition c_TByte ≝ TByte.
+
+coercion c_TByte.
(* introduzione di un tipo dipendente (dalla modalita') per gli argomenti *)
inductive MA_check : instr_mode → Type ≝
inductive instruction : Type ≝
instr: ∀i:instr_mode.opcode → MA_check i → instruction.
-coercion instr.
+definition c_instr ≝ instr.
+
+coercion c_instr.
(* picker: trasforma l'argomento necessario in input a bytes_of_pseudo_instr_mode_param:
MA_check i → list (t_byte8 m) *)
include "nat/minus.ma".
include "list/list.ma".
-interpretation "list nth" 'nth = (nth _).
-interpretation "list nth" 'nth_appl l d i = (nth _ l d i).
+interpretation "list nth" 'nth = (nth ?).
+interpretation "list nth" 'nth_appl l d i = (nth ? l d i).
notation "\nth" with precedence 90 for @{'nth}.
notation < "\nth \nbsp term 90 l \nbsp term 90 d \nbsp term 90 i"
with precedence 69 for @{'nth_appl $l $d $i}.
match n with [ O ⇒ nil ? | S m ⇒ def m :: make_list m]
in make_list.
-interpretation "\mk_list appl" 'mk_list_appl f n = (make_list _ f n).
-interpretation "\mk_list" 'mk_list = (make_list _).
+interpretation "\mk_list appl" 'mk_list_appl f n = (make_list ? f n).
+interpretation "\mk_list" 'mk_list = (make_list ?).
notation "\mk_list" with precedence 90 for @{'mk_list}.
notation < "\mk_list \nbsp term 90 f \nbsp term 90 n"
with precedence 69 for @{'mk_list_appl $f $n}.
notation "\len" with precedence 90 for @{'len}.
-interpretation "len" 'len = (length _).
+interpretation "len" 'len = (length ?).
notation < "\len \nbsp term 90 l" with precedence 69 for @{'len_appl $l}.
-interpretation "len appl" 'len_appl l = (length _ l).
+interpretation "len appl" 'len_appl l = (length ? l).
lemma mk_list_ext: ∀T:Type.∀f1,f2:nat→T.∀n. (∀x.x<n → f1 x = f2 x) → \mk_list f1 n = \mk_list f2 n.
intros 4;elim n; [reflexivity] simplify; rewrite > H1; [2: apply le_n]
notation > "\last" non associative with precedence 90 for @{'last}.
notation < "\last d l" non associative with precedence 70 for @{'last2 $d $l}.
-interpretation "list last" 'last = (last _).
-interpretation "list last applied" 'last2 d l = (last _ d l).
+interpretation "list last" 'last = (last ?).
+interpretation "list last applied" 'last2 d l = (last ? d l).
definition head ≝
λT:Type.λd.λl:list T.\nth l d O.
notation > "\hd" non associative with precedence 90 for @{'hd}.
notation < "\hd d l" non associative with precedence 70 for @{'hd2 $d $l}.
-interpretation "list head" 'hd = (head _).
-interpretation "list head applied" 'hd2 d l = (head _ d l).
+interpretation "list head" 'hd = (head ?).
+interpretation "list head applied" 'hd2 d l = (head ? d l).
definition q2_lt := mk_rel bar (λx,y:bar.\fst x < \fst y).
-interpretation "bar lt" 'lt x y = (rel_op _ q2_lt x y).
+interpretation "bar lt" 'lt x y = (rel_op ? q2_lt x y).
lemma q2_trans : ∀a,b,c:bar. a < b → b < c → a < c.
intros 3; cases a; cases b; cases c; unfold q2_lt; simplify; intros;
coercion canonical_q_lt with nocomposites.
-interpretation "bar lt" 'lt x y = (FunClass_2_OF_trans_rel (canonical_q_lt _) x y).
+interpretation "bar lt" 'lt x y = (FunClass_2_OF_trans_rel (canonical_q_lt ?) x y).
definition nth_base ≝ λf,n. \fst (\nth f ▭ n).
definition nth_height ≝ λf,n. \snd (\nth f ▭ n).
qed.
notation < "x \blacksquare" non associative with precedence 50 for @{'unpos $x}.
-interpretation "hide unpos proof" 'unpos x = (unpos x _).
+interpretation "hide unpos proof" 'unpos x = (unpos x ?).
non associative with precedence 45
for @{ 'equiv $a $b }.
-interpretation "uguaglianza parziale" 'equiv x y = (eq _ x y).
+interpretation "uguaglianza parziale" 'equiv x y = (eq ? x y).
coercion cic:/matita/tests/decl/L.ind#xpointer(1/1/2).
a_algebra_properties: is_algebra ? ? a_mult a_one
}.
-interpretation "Algebra product" 'times a b = (a_mult _ a b).
+interpretation "Algebra product" 'times a b = (a_mult ? a b).
definition ring_of_algebra ≝
λK.λA:algebra K.
apply (not_eq_zero_one_ ? ? ? (r_ring_properties R)).
qed.
-interpretation "Ring mult" 'times a b = (mult _ a b).
+interpretation "Ring mult" 'times a b = (mult ? a b).
notation "1" with precedence 89
for @{ 'one }.
-interpretation "Ring one" 'one = (one _).
+interpretation "Ring one" 'one = (one ?).
lemma eq_mult_zero_x_zero: ∀R:ring.∀x:R.0*x=0.
intros;
vs_vector_space_properties :> is_vector_space ? vs_abelian_group emult
}.
-interpretation "Vector space external product" 'times a b = (emult _ _ a b).
+interpretation "Vector space external product" 'times a b = (emult ? ? a b).
record is_semi_norm (R:real) (V: vector_space R) (semi_norm:V→R) : Prop \def
{ sn_positive: ∀x:V. zero R ≤ semi_norm x;
notation "hvbox(x break ∈ A)" with precedence 60
for @{ 'member_of $x $A }.
-interpretation "Member of" 'member_of x A = (mk_member_of _ A x).
+interpretation "Member of" 'member_of x A = (mk_member_of ? A x).
notation "hvbox(x break ∉ A)" with precedence 60
for @{ 'not_member_of $x $A }.
-interpretation "Not member of" 'not_member_of x A = (Not (member_of _ A x)).
+interpretation "Not member of" 'not_member_of x A = (Not (member_of ? A x)).
definition emptyset : ∀X.set X ≝ λX:Type.λx:X.False.
notation "∅︀" with precedence 100 for @{ 'emptyset }.
-interpretation "Emptyset" 'emptyset = (emptyset _).
+interpretation "Emptyset" 'emptyset = (emptyset ?).
definition subset: ∀X. set X → set X → Prop≝ λX.λA,B:set X.∀x. x ∈ A → x ∈ B.
notation "hvbox(A break ⊆ B)" with precedence 60
for @{ 'subset $A $B }.
-interpretation "Subset" 'subset A B = (subset _ A B).
+interpretation "Subset" 'subset A B = (subset ? A B).
definition intersection: ∀X. set X → set X → set X ≝
λX.λA,B:set X.λx. x ∈ A ∧ x ∈ B.
notation "hvbox(A break ∩ B)" with precedence 70
for @{ 'intersection $A $B }.
-interpretation "Intersection" 'intersection A B = (intersection _ A B).
+interpretation "Intersection" 'intersection A B = (intersection ? A B).
definition union: ∀X. set X → set X → set X ≝ λX.λA,B:set X.λx. x ∈ A ∨ x ∈ B.
notation "hvbox(A break ∪ B)" with precedence 65
for @{ 'union $A $B }.
-interpretation "Union" 'union A B = (union _ A B).
+interpretation "Union" 'union A B = (union ? A B).
definition seq ≝ λX:Type.nat → X.
notation "hvbox(A \sub i)" with precedence 100
for @{ 'nth $A $i }.
-interpretation "nth" 'nth A i = (nth _ A i).
+interpretation "nth" 'nth A i = (nth ? A i).
definition countable_union: ∀X. seq (set X) → set X ≝
λX.λA:seq (set X).λx.∃j.x ∈ A \sub j.
notation "∪ \sub (ident i opt (: ty)) B" with precedence 65
for @{ 'big_union ${default @{(λ${ident i}:$ty.$B)} @{(λ${ident i}.$B)}}}.
-interpretation "countable_union" 'big_union η.t = (countable_union _ t).
+interpretation "countable_union" 'big_union η.t = (countable_union ? t).
definition complement: ∀X. set X \to set X ≝ λX.λA:set X.λx. x ∉ A.
notation "A \sup 'c'" with precedence 100
for @{ 'complement $A }.
-interpretation "Complement" 'complement A = (complement _ A).
+interpretation "Complement" 'complement A = (complement ? A).
definition inverse_image: ∀X,Y.∀f: X → Y.set Y → set X ≝
λX,Y,f,B,x. f x ∈ B.
notation "hvbox(f \sup (-1))" with precedence 100
for @{ 'finverse $f }.
-interpretation "Inverse image" 'finverse f = (inverse_image _ _ f).
+interpretation "Inverse image" 'finverse f = (inverse_image ? ? f).
@{\lambda ${ident i} . $p}}}.
*)
-interpretation "constructive exists" 'exists \eta.x = (ex _ x).
-interpretation "constructive exists (Type)" 'exists \eta.x = (exT _ x).
+interpretation "constructive exists" 'exists \eta.x = (ex ? x).
+interpretation "constructive exists (Type)" 'exists \eta.x = (exT ? x).
alias id "False" = "cic:/matita/logic/connectives/False.ind#xpointer(1/1)".
definition Not ≝ λx:Type.x → False.
let rec gpow (G : abelian_group) (x : G) (n : nat) on n : G ≝
match n with [ O ⇒ 0 | S m ⇒ x + gpow ? x m].
-interpretation "additive abelian group pow" 'times n x = (gpow _ x n).
+interpretation "additive abelian group pow" 'times n x = (gpow ? x n).
record dgroup : Type ≝ {
dg_carr:> abelian_group;
intros (G x n); cases (dg_prop G x n); apply w;
qed.
-interpretation "divisible group divide" 'divide x n = (divide _ x n).
+interpretation "divisible group divide" 'divide x n = (divide ? x n).
lemma divide_divides:
∀G:dgroup.∀x:G.∀n. S n * (x / n) ≈ x.
exc_cotransitive: cotransitive ? exc_excess
}.
-interpretation "Excess base excess" 'nleq a b = (exc_excess _ a b).
+interpretation "Excess base excess" 'nleq a b = (exc_excess ? a b).
(* E(#,≰) → E(#,sym(≰)) *)
lemma make_dual_exc: excess_base → excess_base.
}.
notation "hvbox(a break # b)" non associative with precedence 50 for @{ 'apart $a $b}.
-interpretation "apartness" 'apart x y = (ap_apart _ x y).
+interpretation "apartness" 'apart x y = (ap_apart ? x y).
definition apartness_of_excess_base: excess_base → apartness.
intros (E); apply (mk_apartness E (λa,b:E. a ≰ b ∨ b ≰ a));
definition eq ≝ λA:apartness.λa,b:A. ¬ (a # b).
notation "hvbox(a break ≈ b)" non associative with precedence 50 for @{ 'napart $a $b}.
-interpretation "Apartness alikeness" 'napart a b = (eq _ a b).
-interpretation "Excess alikeness" 'napart a b = (eq (excess_base_OF_excess1 _) a b).
-interpretation "Excess (dual) alikeness" 'napart a b = (eq (excess_base_OF_excess _) a b).
+interpretation "Apartness alikeness" 'napart a b = (eq ? a b).
+interpretation "Excess alikeness" 'napart a b = (eq (excess_base_OF_excess1 ?) a b).
+interpretation "Excess (dual) alikeness" 'napart a b = (eq (excess_base_OF_excess ?) a b).
lemma eq_reflexive:∀E:apartness. reflexive ? (eq E).
intros (E); unfold; intros (x); apply ap_coreflexive;
λE,x,y,z.eq_trans_ E x z y.
notation > "'Eq'≈" non associative with precedence 50 for @{'eqrewrite}.
-interpretation "eq_rew" 'eqrewrite = (eq_trans _ _ _).
+interpretation "eq_rew" 'eqrewrite = (eq_trans ? ? ?).
alias id "antisymmetric" = "cic:/matita/constructive_higher_order_relations/antisymmetric.con".
lemma le_antisymmetric:
definition lt ≝ λE:excess.λa,b:E. a ≤ b ∧ a # b.
-interpretation "ordered sets less than" 'lt a b = (lt _ a b).
+interpretation "ordered sets less than" 'lt a b = (lt ? a b).
lemma lt_coreflexive: ∀E.coreflexive ? (lt E).
intros 2 (E x); intro H; cases H (_ ABS);
qed.
notation > "'Le'≪" non associative with precedence 50 for @{'lerewritel}.
-interpretation "le_rewl" 'lerewritel = (le_rewl _ _ _).
+interpretation "le_rewl" 'lerewritel = (le_rewl ? ? ?).
notation > "'Le'≫" non associative with precedence 50 for @{'lerewriter}.
-interpretation "le_rewr" 'lerewriter = (le_rewr _ _ _).
+interpretation "le_rewr" 'lerewriter = (le_rewr ? ? ?).
lemma ap_rewl: ∀A:apartness.∀x,z,y:A. x ≈ y → y # z → x # z.
intros (A x z y Exy Ayz); cases (ap_cotransitive ???x Ayz); [2:assumption]
qed.
notation > "'Ap'≪" non associative with precedence 50 for @{'aprewritel}.
-interpretation "ap_rewl" 'aprewritel = (ap_rewl _ _ _).
+interpretation "ap_rewl" 'aprewritel = (ap_rewl ? ? ?).
notation > "'Ap'≫" non associative with precedence 50 for @{'aprewriter}.
-interpretation "ap_rewr" 'aprewriter = (ap_rewr _ _ _).
+interpretation "ap_rewr" 'aprewriter = (ap_rewr ? ? ?).
alias symbol "napart" = "Apartness alikeness".
lemma exc_rewl: ∀A:excess.∀x,z,y:A. x ≈ y → y ≰ z → x ≰ z.
qed.
notation > "'Ex'≪" non associative with precedence 50 for @{'excessrewritel}.
-interpretation "exc_rewl" 'excessrewritel = (exc_rewl _ _ _).
+interpretation "exc_rewl" 'excessrewritel = (exc_rewl ? ? ?).
notation > "'Ex'≫" non associative with precedence 50 for @{'excessrewriter}.
-interpretation "exc_rewr" 'excessrewriter = (exc_rewr _ _ _).
+interpretation "exc_rewr" 'excessrewriter = (exc_rewr ? ? ?).
lemma lt_rewr: ∀A:excess.∀x,z,y:A. x ≈ y → z < y → z < x.
intros (A x y z E H); split; elim H;
qed.
notation > "'Lt'≪" non associative with precedence 50 for @{'ltrewritel}.
-interpretation "lt_rewl" 'ltrewritel = (lt_rewl _ _ _).
+interpretation "lt_rewl" 'ltrewritel = (lt_rewl ? ? ?).
notation > "'Lt'≫" non associative with precedence 50 for @{'ltrewriter}.
-interpretation "lt_rewr" 'ltrewriter = (lt_rewr _ _ _).
+interpretation "lt_rewr" 'ltrewriter = (lt_rewr ? ? ?).
lemma lt_le_transitive: ∀A:excess.∀x,y,z:A.x < y → y ≤ z → x < z.
intros (A x y z LT LE); cases LT (LEx APx); split; [apply (le_transitive ???? LEx LE)]
notation "0" with precedence 89 for @{ 'zero }.
-interpretation "Abelian group zero" 'zero = (zero _).
+interpretation "Abelian group zero" 'zero = (zero ?).
-interpretation "Abelian group plus" 'plus a b = (plus _ a b).
+interpretation "Abelian group plus" 'plus a b = (plus ? a b).
-interpretation "Abelian group opp" 'uminus a = (opp _ a).
+interpretation "Abelian group opp" 'uminus a = (opp ? a).
definition minus ≝
λG:abelian_group.λa,b:G. a + -b.
-interpretation "Abelian group minus" 'minus a b = (minus _ a b).
+interpretation "Abelian group minus" 'minus a b = (minus ? a b).
lemma plus_assoc: ∀G:abelian_group.∀x,y,z:G.x+(y+z)≈x+y+z ≝ plus_assoc_.
lemma plus_comm: ∀G:abelian_group.∀x,y:G.x+y≈y+x ≝ plus_comm_.
}.
notation "a \cross b" left associative with precedence 50 for @{ 'op $a $b }.
-interpretation "semi lattice base operation" 'op a b = (sl_op _ a b).
+interpretation "semi lattice base operation" 'op a b = (sl_op ? a b).
lemma excess_of_semi_lattice_base: semi_lattice_base → excess.
intro l;
sl_lem: ∀x,y.(sl_meet x y) ≤ y
}.
-interpretation "semi lattice meet" 'and a b = (sl_meet _ a b).
+interpretation "semi lattice meet" 'and a b = (sl_meet ? a b).
lemma sl_feq_ml: ∀ml:semi_lattice.∀a,b,c:ml. a ≈ b → (c ∧ a) ≈ (c ∧ b).
intros (l a b c H); unfold eq in H ⊢ %; unfold Not in H ⊢ %;
definition hole ≝ λT:Type.λx:T.x.
notation < "\ldots" non associative with precedence 50 for @{'hole}.
-interpretation "hole" 'hole = (hole _ _).
+interpretation "hole" 'hole = (hole ? ?).
(* SL(E,M,H2-5(#),H67(≰)) → E' → c E = c E' → H67'(≰') → SL(E,M,p2-5,H67') *)
lemma subst_excess:
notation "'le_to_eqj'" non associative with precedence 50 for @{'le_to_eqj}.
notation "'lej'" non associative with precedence 50 for @{'lej}.
-interpretation "Lattice meet" 'meet = (sl_meet (latt_mcarr _)).
-interpretation "Lattice meet_refl" 'meet_refl = (sl_meet_refl (latt_mcarr _)).
-interpretation "Lattice meet_comm" 'meet_comm = (sl_meet_comm (latt_mcarr _)).
-interpretation "Lattice meet_assoc" 'meet_assoc = (sl_meet_assoc (latt_mcarr _)).
-interpretation "Lattice strong_extm" 'strong_extm = (sl_strong_extm (latt_mcarr _)).
-interpretation "Lattice le_to_eqm" 'le_to_eqm = (sl_le_to_eqm (latt_mcarr _)).
-interpretation "Lattice lem" 'lem = (sl_lem (latt_mcarr _)).
-interpretation "Lattice join" 'join = (sl_meet (latt_jcarr _)).
-interpretation "Lattice join_refl" 'join_refl = (sl_meet_refl (latt_jcarr _)).
-interpretation "Lattice join_comm" 'join_comm = (sl_meet_comm (latt_jcarr _)).
-interpretation "Lattice join_assoc" 'join_assoc = (sl_meet_assoc (latt_jcarr _)).
-interpretation "Lattice strong_extm" 'strong_extj = (sl_strong_extm (latt_jcarr _)).
-interpretation "Lattice le_to_eqj" 'le_to_eqj = (sl_le_to_eqm (latt_jcarr _)).
-interpretation "Lattice lej" 'lej = (sl_lem (latt_jcarr _)).
+interpretation "Lattice meet" 'meet = (sl_meet (latt_mcarr ?)).
+interpretation "Lattice meet_refl" 'meet_refl = (sl_meet_refl (latt_mcarr ?)).
+interpretation "Lattice meet_comm" 'meet_comm = (sl_meet_comm (latt_mcarr ?)).
+interpretation "Lattice meet_assoc" 'meet_assoc = (sl_meet_assoc (latt_mcarr ?)).
+interpretation "Lattice strong_extm" 'strong_extm = (sl_strong_extm (latt_mcarr ?)).
+interpretation "Lattice le_to_eqm" 'le_to_eqm = (sl_le_to_eqm (latt_mcarr ?)).
+interpretation "Lattice lem" 'lem = (sl_lem (latt_mcarr ?)).
+interpretation "Lattice join" 'join = (sl_meet (latt_jcarr ?)).
+interpretation "Lattice join_refl" 'join_refl = (sl_meet_refl (latt_jcarr ?)).
+interpretation "Lattice join_comm" 'join_comm = (sl_meet_comm (latt_jcarr ?)).
+interpretation "Lattice join_assoc" 'join_assoc = (sl_meet_assoc (latt_jcarr ?)).
+interpretation "Lattice strong_extm" 'strong_extj = (sl_strong_extm (latt_jcarr ?)).
+interpretation "Lattice le_to_eqj" 'le_to_eqj = (sl_le_to_eqm (latt_jcarr ?)).
+interpretation "Lattice lej" 'lej = (sl_lem (latt_jcarr ?)).
notation "'feq_jl'" non associative with precedence 50 for @{'feq_jl}.
notation "'feq_jr'" non associative with precedence 50 for @{'feq_jr}.
notation "'feq_ml'" non associative with precedence 50 for @{'feq_ml}.
notation "'feq_mr'" non associative with precedence 50 for @{'feq_mr}.
-interpretation "Lattice feq_jl" 'feq_jl = (sl_feq_ml (latt_jcarr _)).
-interpretation "Lattice feq_jr" 'feq_jr = (sl_feq_mr (latt_jcarr _)).
-interpretation "Lattice feq_ml" 'feq_ml = (sl_feq_ml (latt_mcarr _)).
-interpretation "Lattice feq_mr" 'feq_mr = (sl_feq_mr (latt_mcarr _)).
+interpretation "Lattice feq_jl" 'feq_jl = (sl_feq_ml (latt_jcarr ?)).
+interpretation "Lattice feq_jr" 'feq_jr = (sl_feq_mr (latt_jcarr ?)).
+interpretation "Lattice feq_ml" 'feq_ml = (sl_feq_ml (latt_mcarr ?)).
+interpretation "Lattice feq_mr" 'feq_mr = (sl_feq_mr (latt_mcarr ?)).
-interpretation "Lattive meet le" 'leq a b = (le (excess_OF_lattice1 _) a b).
+interpretation "Lattive meet le" 'leq a b = (le (excess_OF_lattice1 ?) a b).
interpretation "Lattive join le (aka ge)" 'geq a b =
(le (excess_OF_lattice _) a b).
}.
notation < "\nbsp \delta a \nbsp b" non associative with precedence 80 for @{ 'delta2 $a $b}.
-interpretation "metric" 'delta2 a b = (metric _ _ a b).
+interpretation "metric" 'delta2 a b = (metric ? ? a b).
notation "\delta" non associative with precedence 80 for @{ 'delta }.
-interpretation "metric" 'delta = (metric _ _).
+interpretation "metric" 'delta = (metric ? ?).
lemma apart_of_metric_space: ∀R.metric_space R → apartness.
intros (R ms); apply (mk_apartness ? (λa,b:ms.0 < δ a b)); unfold;
join: pml_carr → pml_carr → pml_carr
}.
-interpretation "valued lattice meet" 'and a b = (meet _ _ a b).
+interpretation "valued lattice meet" 'and a b = (meet ? ? a b).
-interpretation "valued lattice join" 'or a b = (join _ _ a b).
+interpretation "valued lattice join" 'or a b = (join ? ? a b).
record premetric_lattice_props (R : todgroup) (ml : premetric_lattice_ R) : Prop ≝ {
prop1a: ∀a : ml.δ (a ∧ a) a ≈ 0;
meet_join_le: ∀x,y,z. value (meet x y) ≤ value (meet x (join y z))
}.
-interpretation "valued lattice meet" 'and a b = (meet _ _ a b).
+interpretation "valued lattice meet" 'and a b = (meet ? ? a b).
-interpretation "valued lattice join" 'or a b = (join _ _ a b).
+interpretation "valued lattice join" 'or a b = (join ? ? a b).
notation < "\nbsp \mu a" non associative with precedence 80 for @{ 'value2 $a}.
-interpretation "lattice value" 'value2 a = (value _ _ a).
+interpretation "lattice value" 'value2 a = (value ? ? a).
notation "\mu" non associative with precedence 80 for @{ 'value }.
-interpretation "lattice value" 'value = (value _ _).
+interpretation "lattice value" 'value = (value ? ?).
lemma feq_joinr: ∀R.∀L:vlattice R.∀x,y,z:L.
μ x ≈ μ y → μ (z ∧ x) ≈ μ (z ∧ y) → μ (z ∨ x) ≈ μ (z ∨ y).
λR.λms:metric_space R.λs:sequence ms.λk.λn. δ (s n) k.
notation "s ⇝ x" non associative with precedence 50 for @{'tends $s $x}.
-interpretation "tends to" 'tends s x = (tends0 _ (d2s _ _ s x)).
+interpretation "tends to" 'tends s x = (tends0 ? (d2s ? ? s x)).
notation "hvbox(A break ⊆ B)" with precedence 59
for @{ 'subseteq $A $B}.
-interpretation "Subseteq" 'subseteq A B = (eq _ A (comp eps B)).
+interpretation "Subseteq" 'subseteq A B = (eq ? A (comp eps B)).
axiom leq_refl: ∀A. A ⊆ A.
axiom leq_antisym: ∀A,B. A ⊆ B → B ⊆ A → A=B.
notation "ℱ\sub 1 x" non associative with precedence 60 for @{'F1 $x}.
notation > "ℱ_1" non associative with precedence 90 for @{F1 ???}.
-interpretation "F1" 'F1 x = (F1 ___ x).
+interpretation "F1" 'F1 x = (F1 ??? x).
notation "ℱ\sub 2 x" non associative with precedence 60 for @{'F2 $x}.
notation > "ℱ_2" non associative with precedence 90 for @{F2 ???}.
-interpretation "F2" 'F2 x = (F2 ___ x).
+interpretation "F2" 'F2 x = (F2 ??? x).
lemma REW : ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.∀X,Y:Fo ?? F.
arrows2 C2 (F (ℱ_1 X)) (F (ℱ_1 Y)) →
notation "ℳ\sub 1 x" non associative with precedence 60 for @{'Fm1 $x}.
notation > "ℳ_1" non associative with precedence 90 for @{Fm1 ?????}.
-interpretation "Fm1" 'Fm1 x = (Fm1 _____ x).
+interpretation "Fm1" 'Fm1 x = (Fm1 ????? x).
notation "ℳ\sub 2 x" non associative with precedence 60 for @{'Fm2 $x}.
notation > "ℳ_2" non associative with precedence 90 for @{Fm2 ?????}.
-interpretation "Fm2" 'Fm2 x = (Fm2 _____ x).
+interpretation "Fm2" 'Fm2 x = (Fm2 ????? x).
definition Fm :
∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.
rel: arrows1 ? concr form
}.
-interpretation "basic pair relation" 'Vdash2 x y c = (fun21 ___ (rel c) x y).
+interpretation "basic pair relation" 'Vdash2 x y c = (fun21 ??? (rel c) x y).
interpretation "basic pair relation (non applied)" 'Vdash c = (rel c).
alias symbol "eq" = "setoid1 eq".
}.
-interpretation "concrete relation" 'concr_rel r = (concr_rel __ r).
-interpretation "formal relation" 'form_rel r = (form_rel __ r).
+interpretation "concrete relation" 'concr_rel r = (concr_rel ?? r).
+interpretation "formal relation" 'form_rel r = (form_rel ?? r).
definition relation_pair_equality:
∀o1,o2. equivalence_relation1 (relation_pair o1 o2).
| apply (. #‡((†e)‡(†e1))); assumption]]
qed.
-interpretation "fintersects" 'fintersects U V = (fun21 ___ (fintersects _) U V).
+interpretation "fintersects" 'fintersects U V = (fun21 ??? (fintersects ?) U V).
definition fintersectsS:
∀o:BP. binary_morphism1 (Ω \sup (form o)) (Ω \sup (form o)) (Ω \sup (form o)).
| apply (. #‡((†e)‡(†e1))); assumption]]
qed.
-interpretation "fintersectsS" 'fintersects U V = (fun21 ___ (fintersectsS _) U V).
+interpretation "fintersectsS" 'fintersects U V = (fun21 ??? (fintersectsS ?) U V).
definition relS: ∀o: BP. binary_morphism1 (concr o) (Ω \sup (form o)) CPROP.
intros (o); constructor 1;
- [ apply (λx:concr o.λS: Ω \sup (form o).∃y:form o.y ∈ S ∧ x ⊩_o y);
+ [ apply (λx:concr o.λS: Ω \sup (form o).∃y:form o.y ∈ S ∧ x ⊩⎽o y);
| intros; split; intros; cases e2; exists [1,3: apply w]
[ apply (. (#‡e1^-1)‡(e^-1‡#)); assumption
| apply (. (#‡e1)‡(e‡#)); assumption]]
qed.
-interpretation "basic pair relation for subsets" 'Vdash2 x y c = (fun21 (concr _) __ (relS c) x y).
-interpretation "basic pair relation for subsets (non applied)" 'Vdash c = (fun21 ___ (relS c)).
+interpretation "basic pair relation for subsets" 'Vdash2 x y c = (fun21 (concr ?) ?? (relS c) x y).
+interpretation "basic pair relation for subsets (non applied)" 'Vdash c = (fun21 ??? (relS c)).
notation > "hvbox(a break =_\ID b)" non associative with precedence 45
for @{ cic:/matita/logic/equality/eq.ind#xpointer(1/1) ? $a $b }.
-interpretation "ID eq" 'eqID x y = (cic:/matita/logic/equality/eq.ind#xpointer(1/1) _ x y).
+interpretation "ID eq" 'eqID x y = (cic:/matita/logic/equality/eq.ind#xpointer(1/1) ? x y).
record equivalence_relation (A:Type0) : Type1 ≝
{ eq_rel:2> A → A → CProp0;
coercion setoid2_of_setoid1.
prefer coercion Type_OF_setoid2.
prefer coercion Type_OF_setoid.
-prefer coercion Type_OF_setoid1.
+prefer coercion Type_OF_setoid1.
(* we prefer 0 < 1 < 2 *)
record equivalence_relation3 (A:Type3) : Type4 ≝
}.
-interpretation "setoid3 eq" 'eq t x y = (eq_rel3 _ (eq3 t) x y).
-interpretation "setoid2 eq" 'eq t x y = (eq_rel2 _ (eq2 t) x y).
-interpretation "setoid1 eq" 'eq t x y = (eq_rel1 _ (eq1 t) x y).
-interpretation "setoid eq" 'eq t x y = (eq_rel _ (eq t) x y).
+interpretation "setoid3 eq" 'eq t x y = (eq_rel3 ? (eq3 t) x y).
+interpretation "setoid2 eq" 'eq t x y = (eq_rel2 ? (eq2 t) x y).
+interpretation "setoid1 eq" 'eq t x y = (eq_rel1 ? (eq1 t) x y).
+interpretation "setoid eq" 'eq t x y = (eq_rel ? (eq t) x y).
notation > "hvbox(a break =_12 b)" non associative with precedence 45
for @{ eq_rel2 (carr2 (setoid2_of_setoid1 ?)) (eq2 (setoid2_of_setoid1 ?)) $a $b }.
notation > "hvbox(a break =_3 b)" non associative with precedence 45
for @{ eq_rel3 ? (eq3 ?) $a $b }.
-interpretation "setoid3 symmetry" 'invert r = (sym3 ____ r).
-interpretation "setoid2 symmetry" 'invert r = (sym2 ____ r).
-interpretation "setoid1 symmetry" 'invert r = (sym1 ____ r).
-interpretation "setoid symmetry" 'invert r = (sym ____ r).
+interpretation "setoid3 symmetry" 'invert r = (sym3 ???? r).
+interpretation "setoid2 symmetry" 'invert r = (sym2 ???? r).
+interpretation "setoid1 symmetry" 'invert r = (sym1 ???? r).
+interpretation "setoid symmetry" 'invert r = (sym ???? r).
notation ".= r" with precedence 50 for @{'trans $r}.
-interpretation "trans3" 'trans r = (trans3 _____ r).
-interpretation "trans2" 'trans r = (trans2 _____ r).
-interpretation "trans1" 'trans r = (trans1 _____ r).
-interpretation "trans" 'trans r = (trans _____ r).
+interpretation "trans3" 'trans r = (trans3 ????? r).
+interpretation "trans2" 'trans r = (trans2 ????? r).
+interpretation "trans1" 'trans r = (trans1 ????? r).
+interpretation "trans" 'trans r = (trans ????? r).
record unary_morphism (A,B: setoid) : Type0 ≝
{ fun1:1> A → B;
notation "† c" with precedence 90 for @{'prop1 $c }.
notation "l ‡ r" with precedence 90 for @{'prop2 $l $r }.
notation "#" with precedence 90 for @{'refl}.
-interpretation "prop1" 'prop1 c = (prop1 _____ c).
-interpretation "prop11" 'prop1 c = (prop11 _____ c).
-interpretation "prop12" 'prop1 c = (prop12 _____ c).
-interpretation "prop13" 'prop1 c = (prop13 _____ c).
-interpretation "prop2" 'prop2 l r = (prop2 ________ l r).
-interpretation "prop21" 'prop2 l r = (prop21 ________ l r).
-interpretation "prop22" 'prop2 l r = (prop22 ________ l r).
-interpretation "prop23" 'prop2 l r = (prop23 ________ l r).
-interpretation "refl" 'refl = (refl ___).
-interpretation "refl1" 'refl = (refl1 ___).
-interpretation "refl2" 'refl = (refl2 ___).
-interpretation "refl3" 'refl = (refl3 ___).
-
-definition unary_morphism2_of_unary_morphism1: ∀S,T.unary_morphism1 S T → unary_morphism2 S T.
+interpretation "prop1" 'prop1 c = (prop1 ????? c).
+interpretation "prop11" 'prop1 c = (prop11 ????? c).
+interpretation "prop12" 'prop1 c = (prop12 ????? c).
+interpretation "prop13" 'prop1 c = (prop13 ????? c).
+interpretation "prop2" 'prop2 l r = (prop2 ???????? l r).
+interpretation "prop21" 'prop2 l r = (prop21 ???????? l r).
+interpretation "prop22" 'prop2 l r = (prop22 ???????? l r).
+interpretation "prop23" 'prop2 l r = (prop23 ???????? l r).
+interpretation "refl" 'refl = (refl ???).
+interpretation "refl1" 'refl = (refl1 ???).
+interpretation "refl2" 'refl = (refl2 ???).
+interpretation "refl3" 'refl = (refl3 ???).
+
+definition unary_morphism2_of_unary_morphism1:
+ ∀S,T:setoid1.unary_morphism1 S T → unary_morphism2 (setoid2_of_setoid1 S) T.
intros;
constructor 1;
[ apply (fun11 ?? u);
qed.
notation ". r" with precedence 50 for @{'fi $r}.
-interpretation "fi" 'fi r = (fi' __ r).
+interpretation "fi" 'fi r = (fi' ?? r).
definition and_morphism: binary_morphism1 CPROP CPROP CPROP.
constructor 1;
| apply (fi ?? e1 b1)]]
qed.
-interpretation "and_morphism" 'and a b = (fun21 ___ and_morphism a b).
+interpretation "and_morphism" 'and a b = (fun21 ??? and_morphism a b).
definition or_morphism: binary_morphism1 CPROP CPROP CPROP.
constructor 1;
| apply (fi ?? e1 b1)]]
qed.
-interpretation "or_morphism" 'or a b = (fun21 ___ or_morphism a b).
+interpretation "or_morphism" 'or a b = (fun21 ??? or_morphism a b).
definition if_morphism: binary_morphism1 CPROP CPROP CPROP.
constructor 1;
notation "'ASSOC'" with precedence 90 for @{'assoc}.
-interpretation "category2 composition" 'compose x y = (fun22 ___ (comp2 ____) y x).
-interpretation "category2 assoc" 'assoc = (comp_assoc2 ________).
-interpretation "category1 composition" 'compose x y = (fun21 ___ (comp1 ____) y x).
-interpretation "category1 assoc" 'assoc = (comp_assoc1 ________).
-interpretation "category composition" 'compose x y = (fun2 ___ (comp ____) y x).
-interpretation "category assoc" 'assoc = (comp_assoc ________).
+interpretation "category2 composition" 'compose x y = (fun22 ??? (comp2 ????) y x).
+interpretation "category2 assoc" 'assoc = (comp_assoc2 ????????).
+interpretation "category1 composition" 'compose x y = (fun21 ??? (comp1 ????) y x).
+interpretation "category1 assoc" 'assoc = (comp_assoc1 ????????).
+interpretation "category composition" 'compose x y = (fun2 ??? (comp ????) y x).
+interpretation "category assoc" 'assoc = (comp_assoc ????????).
definition category2_of_category1: category1 → category2.
intro;
inductive exT22 (A:Type2) (P:A→CProp2) : CProp2 ≝
ex_introT22: ∀w:A. P w → exT22 A P.
-interpretation "CProp2 exists" 'exists \eta.x = (exT22 _ x).
+interpretation "CProp2 exists" 'exists \eta.x = (exT22 ? x).
definition pi1exT22 ≝ λA,P.λx:exT22 A P.match x with [ex_introT22 x _ ⇒ x].
definition pi2exT22 ≝
λA,P.λx:exT22 A P.match x return λx.P (pi1exT22 ?? x) with [ex_introT22 _ p ⇒ p].
-interpretation "exT22 \fst" 'pi1 = (pi1exT22 _ _).
-interpretation "exT22 \snd" 'pi2 = (pi2exT22 _ _).
-interpretation "exT22 \fst a" 'pi1a x = (pi1exT22 _ _ x).
-interpretation "exT22 \snd a" 'pi2a x = (pi2exT22 _ _ x).
-interpretation "exT22 \fst b" 'pi1b x y = (pi1exT22 _ _ x y).
-interpretation "exT22 \snd b" 'pi2b x y = (pi2exT22 _ _ x y).
+interpretation "exT22 \fst" 'pi1 = (pi1exT22 ? ?).
+interpretation "exT22 \snd" 'pi2 = (pi2exT22 ? ?).
+interpretation "exT22 \fst a" 'pi1a x = (pi1exT22 ? ? x).
+interpretation "exT22 \snd a" 'pi2a x = (pi2exT22 ? ? x).
+interpretation "exT22 \fst b" 'pi1b x y = (pi1exT22 ? ? x y).
+interpretation "exT22 \snd b" 'pi2b x y = (pi2exT22 ? ? x y).
inductive exT (A:Type0) (P:A→CProp0) : CProp0 ≝
ex_introT: ∀w:A. P w → exT A P.
-interpretation "CProp exists" 'exists \eta.x = (exT _ x).
+interpretation "CProp exists" 'exists \eta.x = (exT ? x).
notation "\ll term 19 a, break term 19 b \gg"
with precedence 90 for @{'dependent_pair $a $b}.
interpretation "dependent pair" 'dependent_pair a b =
- (ex_introT _ _ a b).
+ (ex_introT ? ? a b).
definition pi1exT ≝ λA,P.λx:exT A P.match x with [ex_introT x _ ⇒ x].
definition pi2exT ≝
λA,P.λx:exT A P.match x return λx.P (pi1exT ?? x) with [ex_introT _ p ⇒ p].
-interpretation "exT \fst" 'pi1 = (pi1exT _ _).
-interpretation "exT \fst a" 'pi1a x = (pi1exT _ _ x).
-interpretation "exT \fst b" 'pi1b x y = (pi1exT _ _ x y).
-interpretation "exT \snd" 'pi2 = (pi2exT _ _).
-interpretation "exT \snd a" 'pi2a x = (pi2exT _ _ x).
-interpretation "exT \snd b" 'pi2b x y = (pi2exT _ _ x y).
+interpretation "exT \fst" 'pi1 = (pi1exT ? ?).
+interpretation "exT \fst a" 'pi1a x = (pi1exT ? ? x).
+interpretation "exT \fst b" 'pi1b x y = (pi1exT ? ? x y).
+interpretation "exT \snd" 'pi2 = (pi2exT ? ?).
+interpretation "exT \snd a" 'pi2a x = (pi2exT ? ? x).
+interpretation "exT \snd b" 'pi2b x y = (pi2exT ? ? x y).
inductive exT23 (A:Type0) (P:A→CProp0) (Q:A→CProp0) (R:A→A→CProp0) : CProp0 ≝
ex_introT23: ∀w,p:A. P w → Q p → R w p → exT23 A P Q R.
definition pi2exT23 ≝
λA,P,Q,R.λx:exT23 A P Q R.match x with [ex_introT23 _ x _ _ _ ⇒ x].
-interpretation "exT2 \fst" 'pi1 = (pi1exT23 _ _ _ _).
-interpretation "exT2 \snd" 'pi2 = (pi2exT23 _ _ _ _).
-interpretation "exT2 \fst a" 'pi1a x = (pi1exT23 _ _ _ _ x).
-interpretation "exT2 \snd a" 'pi2a x = (pi2exT23 _ _ _ _ x).
-interpretation "exT2 \fst b" 'pi1b x y = (pi1exT23 _ _ _ _ x y).
-interpretation "exT2 \snd b" 'pi2b x y = (pi2exT23 _ _ _ _ x y).
+interpretation "exT2 \fst" 'pi1 = (pi1exT23 ? ? ? ?).
+interpretation "exT2 \snd" 'pi2 = (pi2exT23 ? ? ? ?).
+interpretation "exT2 \fst a" 'pi1a x = (pi1exT23 ? ? ? ? x).
+interpretation "exT2 \snd a" 'pi2a x = (pi2exT23 ? ? ? ? x).
+interpretation "exT2 \fst b" 'pi1b x y = (pi1exT23 ? ? ? ? x y).
+interpretation "exT2 \snd b" 'pi2b x y = (pi2exT23 ? ? ? ? x y).
inductive exT2 (A:Type0) (P,Q:A→CProp0) : CProp0 ≝
ex_introT2: ∀w:A. P w → Q w → exT2 A P Q.
try assumption; [ apply (. #‡H) | apply (. #‡H \sup -1)] assumption]
qed.
-interpretation "downarrow" 'downarrow a = (fun_1 __ (downarrow _) a).
+interpretation "downarrow" 'downarrow a = (fun_1 ?? (downarrow ?) a).
definition ffintersects: ∀S:BTop. binary_morphism1 (Ω \sup S) (Ω \sup S) (Ω \sup S).
intros; constructor 1;
| intros; apply (.= (†H)‡(†H1)); apply refl1]
qed.
-interpretation "ffintersects" 'fintersects U V = (fun1 ___ (ffintersects _) U V).
+interpretation "ffintersects" 'fintersects U V = (fun1 ??? (ffintersects ?) U V).
record formal_topology: Type ≝
{ bt:> BTop;
| intros; apply (.= (†H)‡(†H1)); apply refl1]
qed.
-interpretation "ffintersects'" 'fintersects U V = (fun1 ___ (ffintersects' _) U V).
+interpretation "ffintersects'" 'fintersects U V = (fun1 ??? (ffintersects' ?) U V).
record formal_map (S,T: formal_topology) : Type ≝
{ cr:> continuous_relation_setoid S T;
definition hide ≝ λA:Type.λx:A.x.
-interpretation "hide" 'hide x = (hide _ x).
+interpretation "hide" 'hide x = (hide ? x).
qed.
interpretation "unary morphism comprehension with no proof" 'comprehension T P =
- (mk_unary_morphism T _ P _).
+ (mk_unary_morphism T ? P ?).
interpretation "unary morphism1 comprehension with no proof" 'comprehension T P =
- (mk_unary_morphism1 T _ P _).
+ (mk_unary_morphism1 T ? P ?).
notation > "hvbox({ ident i ∈ s | term 19 p | by })" with precedence 90
for @{ 'comprehension_by $s (λ${ident i}. $p) $by}.
for @{ 'comprehension_by $s (λ${ident i}:$_. $p) $by}.
interpretation "unary morphism comprehension with proof" 'comprehension_by s \eta.f p =
- (mk_unary_morphism s _ f p).
+ (mk_unary_morphism s ? f p).
interpretation "unary morphism1 comprehension with proof" 'comprehension_by s \eta.f p =
- (mk_unary_morphism1 s _ f p).
+ (mk_unary_morphism1 s ? f p).
(* per il set-indexing vedere capitolo BPTools (foundational tools), Sect. 0.3.4 complete
lattices, Definizione 0.9 *)
∀p,q.(∀r.oa_overlap p r → oa_overlap q r) → oa_leq p q
}.
-interpretation "o-algebra leq" 'leq a b = (fun21 ___ (oa_leq _) a b).
+interpretation "o-algebra leq" 'leq a b = (fun21 ??? (oa_leq ?) a b).
notation "hovbox(a mpadded width -150% (>)< b)" non associative with precedence 45
for @{ 'overlap $a $b}.
-interpretation "o-algebra overlap" 'overlap a b = (fun21 ___ (oa_overlap _) a b).
+interpretation "o-algebra overlap" 'overlap a b = (fun21 ??? (oa_overlap ?) a b).
notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (∧) \below (\emsp) \nbsp term 90 p)"
non associative with precedence 50 for @{ 'oa_meet $p }.
notation > "hovbox(∧ f)" non associative with precedence 60
for @{ 'oa_meet $f }.
interpretation "o-algebra meet" 'oa_meet f =
- (fun12 __ (oa_meet __) f).
+ (fun12 ?? (oa_meet ??) f).
interpretation "o-algebra meet with explicit function" 'oa_meet_mk f =
- (fun12 __ (oa_meet __) (mk_unary_morphism _ _ f _)).
+ (fun12 ?? (oa_meet ??) (mk_unary_morphism ?? f ?)).
notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (∨) \below (\emsp) \nbsp term 90 p)"
non associative with precedence 50 for @{ 'oa_join $p }.
notation > "hovbox(∨ f)" non associative with precedence 60
for @{ 'oa_join $f }.
interpretation "o-algebra join" 'oa_join f =
- (fun12 __ (oa_join __) f).
+ (fun12 ?? (oa_join ??) f).
interpretation "o-algebra join with explicit function" 'oa_join_mk f =
- (fun12 __ (oa_join __) (mk_unary_morphism _ _ f _)).
+ (fun12 ?? (oa_join ??) (mk_unary_morphism ?? f ?)).
definition binary_meet : ∀O:OAlgebra. binary_morphism1 O O O.
intros; split;
qed.
interpretation "o-algebra binary meet" 'and a b =
- (fun21 ___ (binary_meet _) a b).
+ (fun21 ??? (binary_meet ?) a b).
prefer coercion Type1_OF_OAlgebra.
qed.
interpretation "o-algebra binary join" 'or a b =
- (fun21 ___ (binary_join _) a b).
+ (fun21 ??? (binary_join ?) a b).
lemma oa_overlap_preservers_meet: ∀O:OAlgebra.∀p,q:O.p >< q → p >< (p ∧ q).
(* next change to avoid universe inconsistency *)
for @{ 'oa_join (mk_unary_morphism BOOL ? (λx__:bool.match x__ with [ true ⇒ $a | false ⇒ $b ]) (IF_THEN_ELSE_p ? $a $b)) }.
interpretation "o-algebra join" 'oa_join f =
- (fun12 __ (oa_join __) f).
+ (fun12 ?? (oa_join ??) f).
interpretation "o-algebra join with explicit function" 'oa_join_mk f =
- (fun12 __ (oa_join __) (mk_unary_morphism _ _ f _)).
+ (fun12 ?? (oa_join ??) (mk_unary_morphism ?? f ?)).
record ORelation (P,Q : OAlgebra) : Type2 ≝ {
or_f_ : carr2 (P ⇒ Q);
notation "r \sup ⎻" non associative with precedence 90 for @{'OR_f_minus $r}.
notation > "r⎻" non associative with precedence 90 for @{'OR_f_minus $r}.
-interpretation "o-relation f⎻*" 'OR_f_minus_star r = (fun12 __ (or_f_minus_star _ _) r).
-interpretation "o-relation f⎻" 'OR_f_minus r = (fun12 __ (or_f_minus _ _) r).
-interpretation "o-relation f*" 'OR_f_star r = (fun12 __ (or_f_star _ _) r).
+interpretation "o-relation f⎻*" 'OR_f_minus_star r = (fun12 ?? (or_f_minus_star ? ?) r).
+interpretation "o-relation f⎻" 'OR_f_minus r = (fun12 ?? (or_f_minus ? ?) r).
+interpretation "o-relation f*" 'OR_f_star r = (fun12 ?? (or_f_star ? ?) r).
definition or_prop1 : ∀P,Q:OAlgebra.∀F:ORelation_setoid P Q.∀p,q.
(F p ≤ q) = (p ≤ F* q).
}.
(* FIX *)
-interpretation "o-concrete relation" 'concr_rel r = (Oconcr_rel __ r).
-interpretation "o-formal relation" 'form_rel r = (Oform_rel __ r).
+interpretation "o-concrete relation" 'concr_rel r = (Oconcr_rel ?? r).
+interpretation "o-formal relation" 'form_rel r = (Oform_rel ?? r).
definition Orelation_pair_equality:
∀o1,o2. equivalence_relation2 (Orelation_pair o1 o2).
| apply (. #‡((†H\sup -1)‡(†H1\sup -1))); assumption]]
qed.
-interpretation "fintersects" 'fintersects U V = (fun1 ___ (fintersects _) U V).
+interpretation "fintersects" 'fintersects U V = (fun1 ??? (fintersects ?) U V).
definition fintersectsS:
∀o:BP. binary_morphism1 (Ω \sup (form o)) (Ω \sup (form o)) (Ω \sup (form o)).
| apply (. #‡((†H\sup -1)‡(†H1\sup -1))); assumption]]
qed.
-interpretation "fintersectsS" 'fintersects U V = (fun1 ___ (fintersectsS _) U V).
+interpretation "fintersectsS" 'fintersects U V = (fun1 ??? (fintersectsS ?) U V).
*)
(*
| apply (. (#‡H1 \sup -1)‡(H \sup -1‡#)); assumption]]
qed.
-interpretation "basic pair relation for subsets" 'Vdash2 x y = (fun1 (concr _) __ (relS _) x y).
-interpretation "basic pair relation for subsets (non applied)" 'Vdash = (fun1 ___ (relS _)).
+interpretation "basic pair relation for subsets" 'Vdash2 x y = (fun1 (concr ?) ?? (relS ?) x y).
+interpretation "basic pair relation for subsets (non applied)" 'Vdash = (fun1 ??? (relS ?)).
*)
notation "□ \sub b" non associative with precedence 90 for @{'box $b}.
-notation > "□_term 90 b" non associative with precedence 90 for @{'box $b}.
-interpretation "Universal image ⊩⎻*" 'box x = (fun12 _ _ (or_f_minus_star _ _) (Orel x)).
+notation > "□⎽term 90 b" non associative with precedence 90 for @{'box $b}.
+interpretation "Universal image ⊩⎻*" 'box x = (fun12 ? ? (or_f_minus_star ? ?) (Orel x)).
notation "◊ \sub b" non associative with precedence 90 for @{'diamond $b}.
-notation > "◊_term 90 b" non associative with precedence 90 for @{'diamond $b}.
-interpretation "Existential image ⊩" 'diamond x = (fun12 _ _ (or_f _ _) (Orel x)).
+notation > "◊⎽term 90 b" non associative with precedence 90 for @{'diamond $b}.
+interpretation "Existential image ⊩" 'diamond x = (fun12 ? ? (or_f ? ?) (Orel x)).
notation "'Rest' \sub b" non associative with precedence 90 for @{'rest $b}.
notation > "'Rest'⎽term 90 b" non associative with precedence 90 for @{'rest $b}.
-interpretation "Universal pre-image ⊩*" 'rest x = (fun12 _ _ (or_f_star _ _) (Orel x)).
+interpretation "Universal pre-image ⊩*" 'rest x = (fun12 ? ? (or_f_star ? ?) (Orel x)).
notation "'Ext' \sub b" non associative with precedence 90 for @{'ext $b}.
notation > "'Ext'⎽term 90 b" non associative with precedence 90 for @{'ext $b}.
-interpretation "Existential pre-image ⊩⎻" 'ext x = (fun12 _ _ (or_f_minus _ _) (Orel x)).
\ No newline at end of file
+interpretation "Existential pre-image ⊩⎻" 'ext x = (fun12 ? ? (or_f_minus ? ?) (Orel x)).
\ No newline at end of file
intro t;
constructor 1;
[ apply (Oform t);
- | apply (□_t ∘ Ext⎽t);
- | apply (◊_t ∘ Rest⎽t);
+ | apply (□⎽t ∘ Ext⎽t);
+ | apply (◊⎽t ∘ Rest⎽t);
| apply hide; intros 2; split; intro;
[ change with ((⊩) \sup ⎻* ((⊩) \sup ⎻ U) ≤ (⊩) \sup ⎻* ((⊩) \sup ⎻ V));
apply (. (#‡(lemma_10_4_a ?? (⊩) V)^-1));
apply (. (or_prop2 : ?) ^ -1);
apply oa_leq_refl; ]]
| apply hide; intros 2; split; intro;
- [ change with (◊_t ((⊩) \sup * U) ≤ ◊_t ((⊩) \sup * V));
+ [ change with (◊⎽t ((⊩) \sup * U) ≤ ◊⎽t ((⊩) \sup * V));
apply (. ((lemma_10_4_b ?? (⊩) U)^-1)‡#);
apply (f_image_monotone ?? (⊩) ? ((⊩)* V));
apply f_star_image_monotone;
apply oa_leq_refl; ]]
| apply hide; intros;
apply (.= (oa_overlap_sym' : ?));
- change with ((◊_t ((⊩)* V) >< (⊩)⎻* ((⊩)⎻ U)) = (U >< (◊_t ((⊩)* V))));
+ change with ((◊⎽t ((⊩)* V) >< (⊩)⎻* ((⊩)⎻ U)) = (U >< (◊⎽t ((⊩)* V))));
apply (.= (or_prop3 ?? (⊩) ((⊩)* V) ?));
apply (.= #‡(lemma_10_3_a : ?));
apply (.= (or_prop3 : ?)^-1);
definition A : ∀b:OBP. unary_morphism1 (Oform b) (Oform b).
intros; constructor 1;
- [ apply (λx.□_b (Ext⎽b x));
+ [ apply (λx.□⎽b (Ext⎽b x));
| intros; apply (†(†e));]
qed.
}.
interpretation "o-concrete space downarrow" 'downarrow x =
- (fun11 __ (Odownarrow _) x).
+ (fun11 ?? (Odownarrow ?) x).
definition Obinary_downarrow :
∀C:Oconcrete_space.binary_morphism1 (Oform C) (Oform C) (Oform C).
apply ((†e)‡(†e1));]
qed.
-interpretation "concrete_space binary ↓" 'fintersects a b = (fun21 _ _ _ (Obinary_downarrow _) a b).
+interpretation "concrete_space binary ↓" 'fintersects a b = (fun21 ? ? ? (Obinary_downarrow ?) a b).
record Oconvergent_relation_pair (CS1,CS2: Oconcrete_space) : Type2 ≝
{ Orp:> arrows2 ? CS1 CS2;
try assumption; [ apply (. #‡H) | apply (. #‡H \sup -1)] assumption]
qed.
-interpretation "downarrow" 'downarrow a = (fun_1 __ (downarrow _) a).
+interpretation "downarrow" 'downarrow a = (fun_1 ?? (downarrow ?) a).
definition ffintersects: ∀S:BTop. binary_morphism1 (Ω \sup S) (Ω \sup S) (Ω \sup S).
intros; constructor 1;
| intros; apply (.= (†H)‡(†H1)); apply refl1]
qed.
-interpretation "ffintersects" 'fintersects U V = (fun1 ___ (ffintersects _) U V).
+interpretation "ffintersects" 'fintersects U V = (fun1 ??? (ffintersects ?) U V).
*)
record formal_topology: Type ≝
- { bt:> BTop;
- converges: ∀U,V: bt. A ? (U ↓ V) = (A ? U ∧ A ? V)
+ { bt:> OBTop;
+ converges: ∀U,V: bt. oA bt (U ↓ V) = (oA ? U ∧ oA ? V)
}.
(*
| intros; apply (.= (†H)‡(†H1)); apply refl1]
qed.
-interpretation "ffintersects'" 'fintersects U V = (fun1 ___ (ffintersects' _) U V).
+interpretation "ffintersects'" 'fintersects U V = (fun1 ??? (ffintersects' ?) U V).
*)
record formal_map (S,T: formal_topology) : Type ≝
{ cr:> continuous_relation_setoid S T;
qed.
notation < "F x" left associative with precedence 60 for @{'map_arrows $F $x}.
-interpretation "map arrows 2" 'map_arrows F x = (fun12 _ _ (map_arrows2 _ _ F _ _) x).
+interpretation "map arrows 2" 'map_arrows F x = (fun12 ? ? (map_arrows2 ? ? F ? ?) x).
definition preserve_sup : ∀S,T.∀ f:Ω \sup S ⇒ Ω \sup T. CProp1.
intros (S T f); apply (∀X:Ω \sup S. (f X) = ?);
notation < "hvbox (x \nbsp \natur term 90 r \nbsp y)" with precedence 45 for @{'satisfy $r $x $y}.
notation > "hvbox (x \natur term 90 r y)" with precedence 45 for @{'satisfy $r $x $y}.
-interpretation "relation applied" 'satisfy r x y = (fun21 ___ (satisfy __ r) x y).
+interpretation "relation applied" 'satisfy r x y = (fun21 ??? (satisfy ?? r) x y).
definition binary_relation_setoid: SET → SET → setoid1.
intros (A B);
qed.
interpretation "subset comprehension" 'comprehension s p =
- (comprehension s (mk_unary_morphism1 __ p _)).
+ (comprehension s (mk_unary_morphism1 ?? p ?)).
definition ext: ∀X,S:REL. binary_morphism1 (arrows1 ? X S) S (Ω \sup X).
apply (λX,S.mk_binary_morphism1 ??? (λr:arrows1 REL X S.λf:S.{x ∈ X | x ♮r f}) ?);
*)
interpretation "lifting singl" 'singl x =
- (fun11 _ (objs2 (POW _)) (singleton _) x).
+ (fun11 ? (objs2 (POW ?)) (singleton ?) x).
theorem POW_full: ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? POW S T g = f).
intros; exists;
interpretation "powerset" 'powerset A = (powerset_setoid1 A).
interpretation "subset construction" 'subset \eta.x =
- (mk_powerset_carrier _ (mk_unary_morphism1 _ CPROP x _)).
+ (mk_powerset_carrier ? (mk_unary_morphism1 ? CPROP x ?)).
definition mem: ∀A. binary_morphism1 A (Ω \sup A) CPROP.
intros;
| apply s1; assumption]]
qed.
-interpretation "mem" 'mem a S = (fun21 ___ (mem _) a S).
+interpretation "mem" 'mem a S = (fun21 ??? (mem ?) a S).
definition subseteq: ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) CPROP.
intros;
apply (transitive_subseteq_operator ???? s s4) ]]
qed.
-interpretation "subseteq" 'subseteq U V = (fun21 ___ (subseteq _) U V).
+interpretation "subseteq" 'subseteq U V = (fun21 ??? (subseteq ?) U V).
| apply (. #‡e1); assumption]]
qed.
-interpretation "overlaps" 'overlaps U V = (fun21 ___ (overlaps _) U V).
+interpretation "overlaps" 'overlaps U V = (fun21 ??? (overlaps ?) U V).
definition intersects:
∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) (Ω \sup A).
| apply (. (#‡e)‡(#‡e1)); assumption]]
qed.
-interpretation "intersects" 'intersects U V = (fun21 ___ (intersects _) U V).
+interpretation "intersects" 'intersects U V = (fun21 ??? (intersects ?) U V).
definition union:
∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) (Ω \sup A).
| apply (. (#‡e)‡(#‡e1)); assumption]]
qed.
-interpretation "union" 'union U V = (fun21 ___ (union _) U V).
+interpretation "union" 'union U V = (fun21 ??? (union ?) U V).
(* qua non riesco a mettere set *)
definition singleton: ∀A:setoid. unary_morphism1 A (Ω \sup A).
[ apply a |4: apply a'] try assumption; apply sym; assumption]
qed.
-interpretation "singleton" 'singl a = (fun11 __ (singleton _) a).
+interpretation "singleton" 'singl a = (fun11 ?? (singleton ?) a).
definition big_intersects:
∀A:SET.∀I:SET.unary_morphism2 (setoid1_of_setoid I ⇒ Ω \sup A) (setoid2_of_setoid1 (Ω \sup A)).
for @{ λ${ident a} : S $A. $P }.
notation "'I'" non associative with precedence 90 for @{'I}.
-interpretation "I" 'I = (I_ _).
+interpretation "I" 'I = (I_ ?).
notation < "'I' \lpar a \rpar" non associative with precedence 90 for @{'I1 $a}.
-interpretation "I a" 'I1 a = (I_ _ a).
+interpretation "I a" 'I1 a = (I_ ? a).
notation "'C'" non associative with precedence 90 for @{'C}.
-interpretation "C" 'C = (C_ _).
+interpretation "C" 'C = (C_ ?).
notation < "'C' \lpar a \rpar" non associative with precedence 90 for @{'C1 $a}.
-interpretation "C a" 'C1 a = (C_ _ a).
+interpretation "C a" 'C1 a = (C_ ? a).
notation < "'C' \lpar a , i \rpar" non associative with precedence 90 for @{'C2 $a $i}.
-interpretation "C a i" 'C2 a i = (C_ _ a i).
+interpretation "C a i" 'C2 a i = (C_ ? a i).
definition in_subset :=
λA:AxiomSet.λa∈A.λU:Ω^A.content A U a.
notation "hvbox(a break εU)" non associative with precedence 50 for
@{'in_subset $a $U}.
-interpretation "In subset" 'in_subset a U = (in_subset _ a U).
+interpretation "In subset" 'in_subset a U = (in_subset ? a U).
definition for_all ≝ λA:AxiomSet.λU:Ω^A.λP:A → CProp.∀x.xεU → P x.
| infinity_ : ∀a.∀i : I a. for_all A (C a i) (covers A U) → covers A U a.
notation "'refl'" non associative with precedence 90 for @{'refl}.
-interpretation "refl" 'refl = (refl_ _ _ _).
+interpretation "refl" 'refl = (refl_ ? ? ?).
notation "'infinity'" non associative with precedence 90 for @{'infinity}.
-interpretation "infinity" 'infinity = (infinity_ _ _ _).
+interpretation "infinity" 'infinity = (infinity_ ? ? ?).
notation "U ⊲ V" non associative with precedence 45
for @{ 'covers_subsets $U $V}.
-interpretation "Covers subsets" 'covers_subsets U V = (for_all _ U (covers _ V)).
-interpretation "Covers elem subset" 'covers_subsets U V = (covers _ V U).
+interpretation "Covers subsets" 'covers_subsets U V = (for_all ? U (covers ? V)).
+interpretation "Covers elem subset" 'covers_subsets U V = (covers ? V U).
definition subseteq := λA:AxiomSet.λU,V:Ω^A.∀x.xεU → xεV.
-interpretation "subseteq" 'subseteq u v = (subseteq _ u v).
+interpretation "subseteq" 'subseteq u v = (subseteq ? u v).
definition covers_elim_ ≝
in
aux.
-interpretation "char" 'subset p = (mk_powerset _ p).
+interpretation "char" 'subset p = (mk_powerset ? p).
definition covers_elim :
∀A:AxiomSet.∀U: Ω^A.∀P:A→CProp.∀H1: U ⊆ {x | P x}.
qed.
notation "'trans'" non associative with precedence 90 for @{'trans}.
-interpretation "trans" 'trans = (trans_ _ _).
+interpretation "trans" 'trans = (trans_ ? ?).
| dall_intro: (∀d. cin ? d → P d) → dall D P
.
-interpretation "internal for all" 'iforall η.x = (dall _ x).
+interpretation "internal for all" 'iforall η.x = (dall ? x).
(* internal existential quantification *)
inductive dex (D:Domain) (P:D → Prop): Prop ≝
| dex_intro: ∀d. cin D d → P d → dex D P
.
-interpretation "internal exists" 'iexists η.x = (dex _ x).
+interpretation "internal exists" 'iexists η.x = (dex ? x).
default "absurd"
cic:/Coq/Init/Logic/absurd.con.
-interpretation "Coq's leibnitz's equality" 'eq x y = (cic:/Coq/Init/Logic/eq.ind#xpointer(1/1) _ x y).
+interpretation "Coq's leibnitz's equality" 'eq x y = (cic:/Coq/Init/Logic/eq.ind#xpointer(1/1) ? x y).
theorem f_equal1 : \forall A,B:Type.\forall f:A\to B.\forall x,y:A.
x = y \to (f y) = (f x).
notation "\ee" with precedence 90 for @{ 'neutral }. (* ⅇ *)
notation > "x ⊩ y" with precedence 45 for @{'Vdash2 $x $y ?}.
-notation > "x ⊩_term 90 c y" with precedence 45 for @{'Vdash2 $x $y $c}.
+notation > "x ⊩⎽ term 90 c y" with precedence 45 for @{'Vdash2 $x $y $c}.
notation "x (⊩ \sub term 90 c) y" with precedence 45 for @{'Vdash2 $x $y $c}.
notation > "⊩ " with precedence 60 for @{'Vdash ?}.
notation "(⊩ \sub term 90 c) " with precedence 60 for @{'Vdash $c}.
<row>
<entry></entry>
<entry>|</entry>
- <entry><emphasis role="bold">_</emphasis></entry>
+ <entry><emphasis role="bold">?</emphasis></entry>
<entry>An implicit parameter</entry>
</row>
<row>
|simplify;rewrite > H1 in H;elim (irrefl_Rlt ? H)]]
qed.*)
-interpretation "real numbers exponent" 'exp a b = (Rexp a b _).
+interpretation "real numbers exponent" 'exp a b = (Rexp a b ?).
(* CSC: multiple interpretations in the same file are not considered in the
right order
-interpretation "Finite_enumerable representation" 'repr C i = (repr C _ i).*)
+interpretation "Finite_enumerable representation" 'repr C i = (repr C ? i).*)
-interpretation "Finite_enumerable order" 'card C = (order C _).
+interpretation "Finite_enumerable order" 'card C = (order C ?).
record finite_enumerable_SemiGroup : Type≝
{ semigroup:> SemiGroup;
interpretation "Index_of_finite_enumerable representation"
'index_of_finite_enumerable_semigroup e
=
- (index_of _ (is_finite_enumerable _) e).
+ (index_of ? (is_finite_enumerable ?) e).
(* several definitions/theorems to be moved somewhere else *)
group_properties:> isGroup pregroup
}.
-interpretation "Group inverse" 'invert x = (inv _ x).
+interpretation "Group inverse" 'invert x = (inv ? x).
definition left_cancellable ≝
λT:Type. λop: T -> T -> T.
for @{ 'subgroupimage $H $x }.
interpretation "Subgroup image" 'subgroupimage H x =
- (image _ _ (morphism_OF_subgroup _ H) x).
+ (image ?? (morphism_OF_subgroup ? H) x).
definition member_of_subgroup ≝
λG.λH:subgroup G.λx:G.∃y.x=y \sub H.
for @{ 'not_member_of $x $H }.
interpretation "Member of subgroup" 'member_of x H =
- (member_of_subgroup _ H x).
+ (member_of_subgroup ? H x).
interpretation "Not member of subgroup" 'not_member_of x H =
- (Not (member_of_subgroup _ H x)).
+ (Not (member_of_subgroup ? H x)).
(* Left cosets *)
(* Here I would prefer 'magma_op, but this breaks something in the next definition *)
interpretation "Left_coset" 'times x C =
- (mk_left_coset _ x C).
+ (mk_left_coset ? x C).
definition member_of_left_coset ≝
λG:Group.λC:left_coset G.λx:G.
∃y.x=(element ? C)·y \sub (subgrp ? C).
interpretation "Member of left_coset" 'member_of x C =
- (member_of_left_coset _ C x).
+ (member_of_left_coset ? C x).
definition left_coset_eq ≝
λG.λC,C':left_coset G.
for @{ 'disjoint $a $b }.
interpretation "Left cosets disjoint" 'disjoint C C' =
- (left_coset_disjoint _ C C').
+ (left_coset_disjoint ? C C').
(* The following should be a one-shot alias! *)
alias symbol "member_of" (instance 0) = "Member of subgroup".
monoid_properties:> isMonoid premonoid
}.
-interpretation "Monoid unit" 'neutral = (e _).
+interpretation "Monoid unit" 'neutral = (e ?).
definition is_left_inverse ≝
λM:Monoid.
op: carrier → carrier → carrier
}.
-interpretation "magma operation" 'middot a b = (op _ a b).
+interpretation "magma operation" 'middot a b = (op ? a b).
(* Semigroups *)
bs_cotransitive: cotransitive ? bs_apart
}.
-interpretation "bishop set apartness" 'apart x y = (bs_apart _ x y).
+interpretation "bishop set apartness" 'apart x y = (bs_apart ? x y).
definition bishop_set_of_ordered_set: ordered_set → bishop_set.
intros (E); apply (mk_bishop_set E (λa,b:E. a ≰ b ∨ b ≰ a));
(* Definition 2.2 (2) *)
definition eq ≝ λA:bishop_set.λa,b:A. ¬ (a # b).
-interpretation "Bishop set alikeness" 'napart a b = (eq _ a b).
+interpretation "Bishop set alikeness" 'napart a b = (eq ? a b).
lemma eq_reflexive:∀E:bishop_set. reflexive ? (eq E).
intros (E); unfold; intros (x); apply bs_coreflexive;
definition bs_subset ≝ λO:bishop_set.λP,Q:O→Prop.∀x:O.P x → Q x.
-interpretation "bishop set subset" 'subseteq a b = (bs_subset _ a b).
+interpretation "bishop set subset" 'subseteq a b = (bs_subset ? a b).
definition square_bishop_set : bishop_set → bishop_set.
intro S; apply (mk_bishop_set (S × S));
notation > "'Eq'≈" non associative with precedence 50
for @{'eqrewrite}.
-interpretation "eq_rew" 'eqrewrite = (eq_trans _ _ _).
+interpretation "eq_rew" 'eqrewrite = (eq_trans ? ? ?).
lemma le_rewl: ∀E:ordered_set.∀z,y,x:E. x ≈ y → x ≤ z → y ≤ z.
intros (E z y x Exy Lxz); apply (le_transitive ??? ? Lxz);
qed.
notation > "'Le'≪" non associative with precedence 50 for @{'lerewritel}.
-interpretation "le_rewl" 'lerewritel = (le_rewl _ _ _).
+interpretation "le_rewl" 'lerewritel = (le_rewl ? ? ?).
notation > "'Le'≫" non associative with precedence 50 for @{'lerewriter}.
-interpretation "le_rewr" 'lerewriter = (le_rewr _ _ _).
+interpretation "le_rewr" 'lerewriter = (le_rewr ? ? ?).
lemma ap_rewl: ∀A:bishop_set.∀x,z,y:A. x ≈ y → y # z → x # z.
intros (A x z y Exy Ayz); cases (bs_cotransitive ???x Ayz); [2:assumption]
qed.
notation > "'Ap'≪" non associative with precedence 50 for @{'aprewritel}.
-interpretation "ap_rewl" 'aprewritel = (ap_rewl _ _ _).
+interpretation "ap_rewl" 'aprewritel = (ap_rewl ? ? ?).
notation > "'Ap'≫" non associative with precedence 50 for @{'aprewriter}.
-interpretation "ap_rewr" 'aprewriter = (ap_rewr _ _ _).
+interpretation "ap_rewr" 'aprewriter = (ap_rewr ? ? ?).
lemma exc_rewl: ∀A:ordered_set.∀x,z,y:A. x ≈ y → y ≰ z → x ≰ z.
intros (A x z y Exy Ayz); cases (exc_cotransitive ?? x Ayz); [2:assumption]
qed.
notation > "'Ex'≪" non associative with precedence 50 for @{'ordered_setrewritel}.
-interpretation "exc_rewl" 'ordered_setrewritel = (exc_rewl _ _ _).
+interpretation "exc_rewl" 'ordered_setrewritel = (exc_rewl ? ? ?).
notation > "'Ex'≫" non associative with precedence 50 for @{'ordered_setrewriter}.
-interpretation "exc_rewr" 'ordered_setrewriter = (exc_rewr _ _ _).
+interpretation "exc_rewr" 'ordered_setrewriter = (exc_rewr ? ? ?).
(*
lemma lt_rewr: ∀A:ordered_set.∀x,z,y:A. x ≈ y → z < y → z < x.
qed.
notation > "'Lt'≪" non associative with precedence 50 for @{'ltrewritel}.
-interpretation "lt_rewl" 'ltrewritel = (lt_rewl _ _ _).
+interpretation "lt_rewl" 'ltrewritel = (lt_rewl ? ? ?).
notation > "'Lt'≫" non associative with precedence 50 for @{'ltrewriter}.
-interpretation "lt_rewr" 'ltrewriter = (lt_rewr _ _ _).
+interpretation "lt_rewr" 'ltrewriter = (lt_rewr ? ? ?).
*)
∀C:ordered_set.
∀a:sequence (os_l C).∀s:segment C.∀H:∀i:nat.a i ∈ s.
∀x:C.∀p:order_converge C a x.
- ∀j. 𝕝_s ≤ (pi1exT23 ???? p j).
+ ∀j. 𝕝_ s ≤ (pi1exT23 ???? p j).
intros; cases p (xi yi Ux Dy Hxy); clear p; simplify;
cases Ux (Ixi Sxi); clear Ux; cases Dy (Dyi Iyi); clear Dy;
cases (Hxy j) (Ia Sa); clear Hxy; cases Ia (Da SSa); cases Sa (Inca SIa); clear Ia Sa;
-intro H2; cases (SSa 𝕝_s H2) (w Hw); simplify in Hw;
+intro H2; cases (SSa 𝕝_ s H2) (w Hw); simplify in Hw;
lapply (H (w+j)) as K; cases (cases_in_segment ? s ? K); apply H3; apply Hw;
qed.
∀C:ordered_set.
∀a:sequence (os_l C).∀s:segment C.∀H:∀i:nat.a i ∈ s.
∀x:C.∀p:order_converge C a x.
- ∀j. (pi2exT23 ???? p j) ≤ 𝕦_s.
+ ∀j. (pi2exT23 ???? p j) ≤ 𝕦_ s.
intros; cases p (xi yi Ux Dy Hxy); clear p; simplify;
cases Ux (Ixi Sxi); clear Ux; cases Dy (Dyi Iyi); clear Dy;
cases (Hxy j) (Ia Sa); clear Hxy; cases Ia (Da SSa); cases Sa (Inca SIa); clear Ia Sa;
-intro H2; cases (SIa 𝕦_s H2) (w Hw); lapply (H (w+j)) as K;
+intro H2; cases (SIa 𝕦_ s H2) (w Hw); lapply (H (w+j)) as K;
cases (cases_in_segment ? s ? K); apply H1; apply Hw;
qed.
cut (∀i.xi i ∈ s) as Hxi; [2:
intros; apply (prove_in_segment (os_l C)); [apply (H3 i)] cases (Hxy i) (H5 _); cases H5 (H7 _);
lapply (H7 0) as K; cases (cases_in_segment ? s ? (H1 i)) (Pl Pu);
- simplify in K:(? ? % ?); apply (hle_transitive (os_l C) (xi i) (a i) 𝕦_s K Pu);] clear H3;
+ simplify in K:(? ? % ?); apply (hle_transitive (os_l C) (xi i) (a i) 𝕦_ s K Pu);] clear H3;
cut (∀i.yi i ∈ s) as Hyi; [2:
intros; apply (prove_in_segment (os_l C)); [2:apply (H2 i)] cases (Hxy i) (_ H5); cases H5 (H7 _);
lapply (H7 0) as K; cases (cases_in_segment ? s ? (H1 i)) (Pl Pu); simplify in K;
- apply (le_transitive 𝕝_s ? ? ? K);apply Pl;] clear H2;
+ apply (le_transitive 𝕝_ s ? ? ? K);apply Pl;] clear H2;
split;
[1: apply (uparrow_to_in_segment s ? Hxi ? Hx);
|2: intros 3 (h);
cut (∀i.xi i ∈ s) as Hxi; [2:
intros; apply (prove_in_segment (os_l C)); [apply (H3 i)] cases (Hxy i) (H5 _); cases H5 (H7 _);
lapply (H7 0) as K; cases (cases_in_segment ? s ? (H1 i)) (Pl Pu);
- simplify in K:(? ? % ?); apply (hle_transitive (os_l C) (xi i) (a i) 𝕦_s K Pu);] clear H3;
+ simplify in K:(? ? % ?); apply (hle_transitive (os_l C) (xi i) (a i) 𝕦_ s K Pu);] clear H3;
cut (∀i.yi i ∈ s) as Hyi; [2:
intros; apply (prove_in_segment (os_l C)); [2:apply (H2 i)] cases (Hxy i) (_ H5); cases H5 (H7 _);
lapply (H7 0) as K; cases (cases_in_segment ? s ? (H1 i)) (Pl Pu); simplify in K;
- apply (le_transitive 𝕝_s ? ? ? K);apply Pl;] clear H2;
+ apply (le_transitive 𝕝_ s ? ? ? K);apply Pl;] clear H2;
letin Xi ≝ (⌊n,≪xi n, Hxi n≫⌋);
letin Yi ≝ (⌊n,≪yi n, Hyi n≫⌋);
cases (restrict_uniform_convergence_uparrow ? S ? (H s) Xi x Hx);
intros 4; cases X (x Hx); clear X; letin X ≝ ≪x,Hx≫;
fold normalize X; intros; cases H1;
alias symbol "N" = "Natural numbers".
-letin spec ≝ (λi,j:ℕ.(𝕦_sg ≤ i ∧ x = \fst (a j)) ∨ (i < 𝕦_sg ∧ x + i ≤ 𝕦_sg + \fst (a j)));
+letin spec ≝ (λi,j:ℕ.(𝕦_ sg ≤ i ∧ x = \fst (a j)) ∨ (i < 𝕦_ sg ∧ x + i ≤ 𝕦_ sg + \fst (a j)));
(* x - aj <= max 0 (u - i) *)
letin m ≝ (hide ? (
let rec aux i ≝
∀i:nat.∃j:nat.spec i j));[whd; apply nP;] unfold spec in aux ⊢ %;
[3: unfold X in H2; clear H4 n aux spec H3 H1 H X;
cases (cases_in_segment ??? Hx);
- elim 𝕦_sg in H1 ⊢ %; intros (a Hs H);
+ elim 𝕦_ sg in H1 ⊢ %; intros (a Hs H);
[1: left; split; [apply le_n]
generalize in match H;
generalize in match Hx;
cases H5; clear H5; cases H7; clear H7;
[1: left; split; [ apply (le_S ?? H5); | assumption]
|3: cases (?:False); rewrite < H8 in H6; apply (not_le_Sn_n ? H6);
- |*: cases (cmp_nat 𝕦_sg (S n1));
+ |*: cases (cmp_nat 𝕦_ sg (S n1));
[1,3: left; split; [1,3: assumption |2: assumption]
- cut (𝕦_sg = S n1); [2: apply le_to_le_to_eq; assumption ]
+ cut (𝕦_ sg = S n1); [2: apply le_to_le_to_eq; assumption ]
clear H7 H5 H4;rewrite > Hcut in H8:(? ? (? % ?)); clear Hcut;
cut (x = S (\fst (a w)));
[2: apply le_to_le_to_eq; [2: assumption]
[1: rewrite > sym_plus in ⊢ (? ? %);
rewrite < H6; apply le_plus_r; assumption;
|2: cases (H3 (a w) H6);
- change with (x + S n1 ≤ 𝕦_sg + \fst (a w1));rewrite < plus_n_Sm;
+ change with (x + S n1 ≤ 𝕦_ sg + \fst (a w1));rewrite < plus_n_Sm;
apply (trans_le ??? (le_S_S ?? H8)); rewrite > plus_n_Sm;
apply (le_plus ???? (le_n ?) H9);]]]]
clearbody m; unfold spec in m; clear spec;
| false ⇒ find (S i) w]]
in find
:
- ∀i,bound.∃j.i + bound = 𝕦_sg → x = \fst (a j));
+ ∀i,bound.∃j.i + bound = 𝕦_ sg → x = \fst (a j));
[1: cases (find (S n) n2); intro; change with (x = \fst (a w));
apply H6; rewrite < H7; simplify; apply plus_n_Sm;
|2: intros; rewrite > (eqb_true_to_eq ?? H5); reflexivity
|3: intros; rewrite > sym_plus in H5; rewrite > H5; clear H5 H4 n n1;
- cases (m 𝕦_sg); cases H4; clear H4; cases H5; clear H5; [assumption]
+ cases (m 𝕦_ sg); cases H4; clear H4; cases H5; clear H5; [assumption]
cases (not_le_Sn_n ? H4)]
-clearbody find; cases (find O 𝕦_sg);
+clearbody find; cases (find O 𝕦_ sg);
exists [apply w]; intros; change with (x = \fst (a j));
rewrite > (H4 ?); [2: reflexivity]
apply le_to_le_to_eq;
notation "''" non associative with precedence 90 for @{'}.
notation "''" non associative with precedence 90 for @{'}.
-interpretation "" ' = ( (os_l _)).
-interpretation "" ' = ( (os_r _)).
+interpretation "" ' = ( (os_l ?)).
+interpretation "" ' = ( (os_r ?)).
*)
(* Definition 2.1 *)
coercion Type_of_ordered_set.
notation "a ≰≰ b" non associative with precedence 45 for @{'nleq_low $a $b}.
-interpretation "Ordered half set excess" 'nleq_low a b = (hos_excess _ a b).
+interpretation "Ordered half set excess" 'nleq_low a b = (hos_excess ? a b).
-interpretation "Ordered set excess (dual)" 'ngeq a b = (hos_excess (os_r _) a b).
-interpretation "Ordered set excess" 'nleq a b = (hos_excess (os_l _) a b).
+interpretation "Ordered set excess (dual)" 'ngeq a b = (hos_excess (os_r ?) a b).
+interpretation "Ordered set excess" 'nleq a b = (hos_excess (os_l ?) a b).
notation "'exc_coreflexive'" non associative with precedence 90 for @{'exc_coreflexive}.
notation "'cxe_coreflexive'" non associative with precedence 90 for @{'cxe_coreflexive}.
-interpretation "exc_coreflexive" 'exc_coreflexive = ((hos_coreflexive (os_l _))).
-interpretation "cxe_coreflexive" 'cxe_coreflexive = ((hos_coreflexive (os_r _))).
+interpretation "exc_coreflexive" 'exc_coreflexive = ((hos_coreflexive (os_l ?))).
+interpretation "cxe_coreflexive" 'cxe_coreflexive = ((hos_coreflexive (os_r ?))).
notation "'exc_cotransitive'" non associative with precedence 90 for @{'exc_cotransitive}.
notation "'cxe_cotransitive'" non associative with precedence 90 for @{'cxe_cotransitive}.
-interpretation "exc_cotransitive" 'exc_cotransitive = ((hos_cotransitive (os_l _))).
-interpretation "cxe_cotransitive" 'cxe_cotransitive = ((hos_cotransitive (os_r _))).
+interpretation "exc_cotransitive" 'exc_cotransitive = ((hos_cotransitive (os_l ?))).
+interpretation "cxe_cotransitive" 'cxe_cotransitive = ((hos_cotransitive (os_r ?))).
(* Definition 2.2 (3) *)
definition le ≝ λE:half_ordered_set.λa,b:E. ¬ (a ≰≰ b).
notation "hvbox(a break ≤≤ b)" non associative with precedence 45 for @{ 'leq_low $a $b }.
-interpretation "Half ordered set greater or equal than" 'leq_low a b = ((le _ a b)).
+interpretation "Half ordered set greater or equal than" 'leq_low a b = ((le ? a b)).
-interpretation "Ordered set greater or equal than" 'geq a b = ((le (os_r _) a b)).
-interpretation "Ordered set less or equal than" 'leq a b = ((le (os_l _) a b)).
+interpretation "Ordered set greater or equal than" 'geq a b = ((le (os_r ?) a b)).
+interpretation "Ordered set less or equal than" 'leq a b = ((le (os_l ?) a b)).
lemma hle_reflexive: ∀E.reflexive ? (le E).
unfold reflexive; intros 3; apply (hos_coreflexive ? x H);
notation "'le_reflexive'" non associative with precedence 90 for @{'le_reflexive}.
notation "'ge_reflexive'" non associative with precedence 90 for @{'ge_reflexive}.
-interpretation "le reflexive" 'le_reflexive = (hle_reflexive (os_l _)).
-interpretation "ge reflexive" 'ge_reflexive = (hle_reflexive (os_r _)).
+interpretation "le reflexive" 'le_reflexive = (hle_reflexive (os_l ?)).
+interpretation "ge reflexive" 'ge_reflexive = (hle_reflexive (os_r ?)).
(* DUALITY TESTS
lemma test_le_ge_convertible :∀o:ordered_set.∀x,y:o. x ≤ y → y ≥ x.
notation "'le_transitive'" non associative with precedence 90 for @{'le_transitive}.
notation "'ge_transitive'" non associative with precedence 90 for @{'ge_transitive}.
-interpretation "le transitive" 'le_transitive = (hle_transitive (os_l _)).
-interpretation "ge transitive" 'ge_transitive = (hle_transitive (os_r _)).
+interpretation "le transitive" 'le_transitive = (hle_transitive (os_l ?)).
+interpretation "ge transitive" 'ge_transitive = (hle_transitive (os_r ?)).
(* Lemma 2.3 *)
lemma exc_hle_variance:
notation "'exc_le_variance'" non associative with precedence 90 for @{'exc_le_variance}.
notation "'exc_ge_variance'" non associative with precedence 90 for @{'exc_ge_variance}.
-interpretation "exc_le_variance" 'exc_le_variance = (exc_hle_variance (os_l _)).
-interpretation "exc_ge_variance" 'exc_ge_variance = (exc_hle_variance (os_r _)).
+interpretation "exc_le_variance" 'exc_le_variance = (exc_hle_variance (os_l ?)).
+interpretation "exc_ge_variance" 'exc_ge_variance = (exc_hle_variance (os_r ?)).
definition square_exc ≝
λO:half_ordered_set.λx,y:O×O.\fst x ≰≰ \fst y ∨ \snd x ≰≰ \snd y.
definition os_subset ≝ λO:ordered_set.λP,Q:O→Prop.∀x:O.P x → Q x.
-interpretation "ordered set subset" 'subseteq a b = (os_subset _ a b).
+interpretation "ordered set subset" 'subseteq a b = (os_subset ? a b).
λC:ordered_set.λU:C squareO → Prop.
λx:C squareO. U 〈\snd x,\fst x〉.
-interpretation "relation invertion" 'invert a = (invert_os_relation _ a).
-interpretation "relation invertion" 'invert_symbol = (invert_os_relation _).
-interpretation "relation invertion" 'invert_appl a x = (invert_os_relation _ a x).
+interpretation "relation invertion" 'invert a = (invert_os_relation ? a).
+interpretation "relation invertion" 'invert_symbol = (invert_os_relation ?).
+interpretation "relation invertion" 'invert_appl a x = (invert_os_relation ? a x).
lemma hint_segment: ∀O.
segment (Type_of_ordered_set O) →
qed.
interpretation "Ordered uniform space segment" 'segset a =
- (segment_ordered_uniform_space _ a).
+ (segment_ordered_uniform_space ? a).
(* Lemma 3.2 *)
alias symbol "pi1" = "exT \fst".
notation "'segment_upperbound'" non associative with precedence 90 for @{'segment_upperbound}.
notation "'segment_lowerbound'" non associative with precedence 90 for @{'segment_lowerbound}.
-interpretation "segment_upperbound" 'segment_upperbound = (h_segment_upperbound (os_l _)).
-interpretation "segment_lowerbound" 'segment_lowerbound = (h_segment_upperbound (os_r _)).
+interpretation "segment_upperbound" 'segment_upperbound = (h_segment_upperbound (os_l ?)).
+interpretation "segment_lowerbound" 'segment_lowerbound = (h_segment_upperbound (os_r ?)).
lemma h_segment_preserves_uparrow:
∀C:half_ordered_set.∀s:segment C.∀a:sequence (half_segment_ordered_set C s).
notation "'segment_preserves_uparrow'" non associative with precedence 90 for @{'segment_preserves_uparrow}.
notation "'segment_preserves_downarrow'" non associative with precedence 90 for @{'segment_preserves_downarrow}.
-interpretation "segment_preserves_uparrow" 'segment_preserves_uparrow = (h_segment_preserves_uparrow (os_l _)).
-interpretation "segment_preserves_downarrow" 'segment_preserves_downarrow = (h_segment_preserves_uparrow (os_r _)).
+interpretation "segment_preserves_uparrow" 'segment_preserves_uparrow = (h_segment_preserves_uparrow (os_l ?)).
+interpretation "segment_preserves_downarrow" 'segment_preserves_downarrow = (h_segment_preserves_uparrow (os_r ?)).
(* Fact 2.18 *)
lemma segment_cauchy:
notation "'uparrow_to_in_segment'" non associative with precedence 90 for @{'uparrow_to_in_segment}.
notation "'downarrow_to_in_segment'" non associative with precedence 90 for @{'downarrow_to_in_segment}.
-interpretation "uparrow_to_in_segment" 'uparrow_to_in_segment = (h_uparrow_to_in_segment (os_l _)).
-interpretation "downarrow_to_in_segment" 'downarrow_to_in_segment = (h_uparrow_to_in_segment (os_r _)).
+interpretation "uparrow_to_in_segment" 'uparrow_to_in_segment = (h_uparrow_to_in_segment (os_l ?)).
+interpretation "downarrow_to_in_segment" 'downarrow_to_in_segment = (h_uparrow_to_in_segment (os_r ?)).
alias symbol "dependent_pair" = "dependent pair".
(* Lemma 3.8 NON DUALIZZATO *)
definition hide ≝ λT:Type.λx:T.x.
notation < "\blacksquare" non associative with precedence 50 for @{'hide}.
-interpretation "hide" 'hide = (hide _ _).
-interpretation "hide2" 'hide = (hide _ _ _).
+interpretation "hide" 'hide = (hide ? ?).
+interpretation "hide2" 'hide = (hide ? ? ?).
definition inject ≝ λP.λa:nat.λp:P a. ex_introT ? P ? p.
coercion inject 0 1.
notation "a \sub i" left associative with precedence 90
for @{ 'sequence_appl $a $i }.
-interpretation "sequence" 'sequence \eta.x = (mk_seq _ x).
-interpretation "sequence element" 'sequence_appl s i = (fun_of_seq _ s i).
+interpretation "sequence" 'sequence \eta.x = (mk_seq ? x).
+interpretation "sequence element" 'sequence_appl s i = (fun_of_seq ? s i).
notation > "x 'is_infimum' s" non associative with precedence 45
for @{'infimum $s $x}.
-interpretation "Ordered set upper bound" 'upper_bound s x = (upper_bound (os_l _) s x).
-interpretation "Ordered set lower bound" 'lower_bound s x = (upper_bound (os_r _) s x).
+interpretation "Ordered set upper bound" 'upper_bound s x = (upper_bound (os_l ?) s x).
+interpretation "Ordered set lower bound" 'lower_bound s x = (upper_bound (os_r ?) s x).
-interpretation "Ordered set increasing" 'increasing s = (increasing (os_l _) s).
-interpretation "Ordered set decreasing" 'decreasing s = (increasing (os_r _) s).
+interpretation "Ordered set increasing" 'increasing s = (increasing (os_l ?) s).
+interpretation "Ordered set decreasing" 'decreasing s = (increasing (os_r ?) s).
-interpretation "Ordered set strong sup" 'supremum s x = (supremum (os_l _) s x).
-interpretation "Ordered set strong inf" 'infimum s x = (supremum (os_r _) s x).
+interpretation "Ordered set strong sup" 'supremum s x = (supremum (os_l ?) s x).
+interpretation "Ordered set strong inf" 'infimum s x = (supremum (os_r ?) s x).
(* Fact 2.5 *)
lemma h_supremum_is_upper_bound:
notation "'supremum_is_upper_bound'" non associative with precedence 90 for @{'supremum_is_upper_bound}.
notation "'infimum_is_lower_bound'" non associative with precedence 90 for @{'infimum_is_lower_bound}.
-interpretation "supremum_is_upper_bound" 'supremum_is_upper_bound = (h_supremum_is_upper_bound (os_l _)).
-interpretation "infimum_is_lower_bound" 'infimum_is_lower_bound = (h_supremum_is_upper_bound (os_r _)).
+interpretation "supremum_is_upper_bound" 'supremum_is_upper_bound = (h_supremum_is_upper_bound (os_l ?)).
+interpretation "infimum_is_lower_bound" 'infimum_is_lower_bound = (h_supremum_is_upper_bound (os_r ?)).
(* Lemma 2.6 *)
definition strictly_increasing ≝
notation > "s 'is_strictly_increasing'" non associative with precedence 45
for @{'strictly_increasing $s}.
interpretation "Ordered set strict increasing" 'strictly_increasing s =
- (strictly_increasing (os_l _) s).
+ (strictly_increasing (os_l ?) s).
notation < "s \nbsp 'is_strictly_decreasing'" non associative with precedence 45
for @{'strictly_decreasing $s}.
notation > "s 'is_strictly_decreasing'" non associative with precedence 45
for @{'strictly_decreasing $s}.
interpretation "Ordered set strict decreasing" 'strictly_decreasing s =
- (strictly_increasing (os_r _) s).
+ (strictly_increasing (os_r ?) s).
definition uparrow ≝
λC:half_ordered_set.λs:sequence C.λu:C.
increasing ? s ∧ supremum ? s u.
-interpretation "Ordered set uparrow" 'funion s u = (uparrow (os_l _) s u).
-interpretation "Ordered set downarrow" 'fintersects s u = (uparrow (os_r _) s u).
+interpretation "Ordered set uparrow" 'funion s u = (uparrow (os_l ?) s u).
+interpretation "Ordered set downarrow" 'fintersects s u = (uparrow (os_r ?) s u).
lemma h_trans_increasing:
∀C:half_ordered_set.∀a:sequence C.increasing ? a →
notation "'trans_increasing'" non associative with precedence 90 for @{'trans_increasing}.
notation "'trans_decreasing'" non associative with precedence 90 for @{'trans_decreasing}.
-interpretation "trans_increasing" 'trans_increasing = (h_trans_increasing (os_l _)).
-interpretation "trans_decreasing" 'trans_decreasing = (h_trans_increasing (os_r _)).
+interpretation "trans_increasing" 'trans_increasing = (h_trans_increasing (os_l ?)).
+interpretation "trans_decreasing" 'trans_decreasing = (h_trans_increasing (os_r ?)).
lemma hint_nat :
Type_OF_ordered_set nat_ordered_set →
notation "'trans_increasing_exc'" non associative with precedence 90 for @{'trans_increasing_exc}.
notation "'trans_decreasing_exc'" non associative with precedence 90 for @{'trans_decreasing_exc}.
-interpretation "trans_increasing_exc" 'trans_increasing_exc = (h_trans_increasing_exc (os_l _)).
-interpretation "trans_decreasing_exc" 'trans_decreasing_exc = (h_trans_increasing_exc (os_r _)).
+interpretation "trans_increasing_exc" 'trans_increasing_exc = (h_trans_increasing_exc (os_l ?)).
+interpretation "trans_decreasing_exc" 'trans_decreasing_exc = (h_trans_increasing_exc (os_r ?)).
alias symbol "exists" = "CProp exists".
lemma nat_strictly_increasing_reaches:
notation "'selection_uparrow'" non associative with precedence 90 for @{'selection_uparrow}.
notation "'selection_downarrow'" non associative with precedence 90 for @{'selection_downarrow}.
-interpretation "selection_uparrow" 'selection_uparrow = (h_selection_uparrow (os_l _)).
-interpretation "selection_downarrow" 'selection_downarrow = (h_selection_uparrow (os_r _)).
+interpretation "selection_uparrow" 'selection_uparrow = (h_selection_uparrow (os_l ?)).
+interpretation "selection_downarrow" 'selection_downarrow = (h_selection_uparrow (os_r ?)).
(* Definition 2.7 *)
definition order_converge ≝
for @{'order_converge $a $x}.
notation > "a 'order_converges' x" non associative with precedence 45
for @{'order_converge $a $x}.
-interpretation "Order convergence" 'order_converge s u = (order_converge _ s u).
+interpretation "Order convergence" 'order_converge s u = (order_converge ? s u).
(* Definition 2.8 *)
record segment (O : Type) : Type ≝ {
seg_u_ : O
}.
-notation > "𝕦_term 90 s" non associative with precedence 90 for @{'upp $s}.
+notation > "𝕦_ term 90 s" non associative with precedence 90 for @{'upp $s}.
notation "𝕦 \sub term 90 s" non associative with precedence 90 for @{'upp $s}.
-notation > "𝕝_term 90 s" non associative with precedence 90 for @{'low $s}.
+notation > "𝕝_ term 90 s" non associative with precedence 90 for @{'low $s}.
notation "𝕝 \sub term 90 s" non associative with precedence 90 for @{'low $s}.
definition seg_u ≝
λO:half_ordered_set.λs:segment O.
wloss O ?? (λl,u.l) (seg_l_ ? s) (seg_u_ ? s).
-interpretation "uppper" 'upp s = (seg_u (os_l _) s).
-interpretation "lower" 'low s = (seg_l (os_l _) s).
-interpretation "uppper dual" 'upp s = (seg_l (os_r _) s).
-interpretation "lower dual" 'low s = (seg_u (os_r _) s).
+interpretation "uppper" 'upp s = (seg_u (os_l ?) s).
+interpretation "lower" 'low s = (seg_l (os_l ?) s).
+interpretation "uppper dual" 'upp s = (seg_l (os_r ?) s).
+interpretation "lower dual" 'low s = (seg_u (os_r ?) s).
definition in_segment ≝
λO:half_ordered_set.λs:segment O.λx:O.
notation "‡O" non associative with precedence 90 for @{'segment $O}.
interpretation "Ordered set sergment" 'segment x = (segment x).
-interpretation "Ordered set sergment in" 'mem x s = (in_segment _ s x).
+interpretation "Ordered set sergment in" 'mem x s = (in_segment ? s x).
definition segment_ordered_set_carr ≝
λO:half_ordered_set.λs:‡O.∃x.x ∈ s.
qed.
notation "{[ term 19 s ]}" non associative with precedence 90 for @{'segset $s}.
-interpretation "Ordered set segment" 'segset s = (segment_ordered_set _ s).
+interpretation "Ordered set segment" 'segset s = (segment_ordered_set ? s).
(* test :
∀O:ordered_set.∀s: segment (os_l O).∀x:O.
notation "'segment_preserves_supremum'" non associative with precedence 90 for @{'segment_preserves_supremum}.
notation "'segment_preserves_infimum'" non associative with precedence 90 for @{'segment_preserves_infimum}.
-interpretation "segment_preserves_supremum" 'segment_preserves_supremum = (h_segment_preserves_supremum (os_l _)).
-interpretation "segment_preserves_infimum" 'segment_preserves_infimum = (h_segment_preserves_supremum (os_r _)).
+interpretation "segment_preserves_supremum" 'segment_preserves_supremum = (h_segment_preserves_supremum (os_l ?)).
+interpretation "segment_preserves_infimum" 'segment_preserves_infimum = (h_segment_preserves_supremum (os_r ?)).
(*
test segment_preserves_infimum2:
notation > "s 'is_upper_located'" non associative with precedence 45
for @{'upper_located $s}.
interpretation "Ordered set upper locatedness" 'upper_located s =
- (upper_located (os_l _) s).
+ (upper_located (os_l ?) s).
notation < "s \nbsp 'is_lower_located'" non associative with precedence 45
for @{'lower_located $s}.
notation > "s 'is_lower_located'" non associative with precedence 45
for @{'lower_located $s}.
interpretation "Ordered set lower locatedness" 'lower_located s =
- (upper_located (os_r _) s).
+ (upper_located (os_r ?) s).
(* Lemma 2.12 *)
lemma h_uparrow_upperlocated:
notation "'uparrow_upperlocated'" non associative with precedence 90 for @{'uparrow_upperlocated}.
notation "'downarrow_lowerlocated'" non associative with precedence 90 for @{'downarrow_lowerlocated}.
-interpretation "uparrow_upperlocated" 'uparrow_upperlocated = (h_uparrow_upperlocated (os_l _)).
-interpretation "downarrow_lowerlocated" 'downarrow_lowerlocated = (h_uparrow_upperlocated (os_r _)).
+interpretation "uparrow_upperlocated" 'uparrow_upperlocated = (h_uparrow_upperlocated (os_l ?)).
+interpretation "downarrow_lowerlocated" 'downarrow_lowerlocated = (h_uparrow_upperlocated (os_r ?)).
λC:ordered_set.λU,V:C squareB → Prop.
λx:C squareB.∃y:C. U 〈\fst x,y〉 ∧ V 〈y,\snd x〉.
-interpretation "bishop set relations composition" 'compose a b = (compose_bs_relations _ a b).
-interpretation "ordered set relations composition" 'compose a b = (compose_os_relations _ a b).
+interpretation "bishop set relations composition" 'compose a b = (compose_bs_relations ? a b).
+interpretation "ordered set relations composition" 'compose a b = (compose_os_relations ? a b).
definition invert_bs_relation ≝
λC:bishop_set.λU:C squareB → Prop.
λx:C squareB. U 〈\snd x,\fst x〉.
notation > "\inv" with precedence 60 for @{ 'invert_symbol }.
-interpretation "relation invertion" 'invert a = (invert_bs_relation _ a).
-interpretation "relation invertion" 'invert_symbol = (invert_bs_relation _).
-interpretation "relation invertion" 'invert_appl a x = (invert_bs_relation _ a x).
+interpretation "relation invertion" 'invert a = (invert_bs_relation ? a).
+interpretation "relation invertion" 'invert_symbol = (invert_bs_relation ?).
+interpretation "relation invertion" 'invert_appl a x = (invert_bs_relation ? a x).
alias symbol "exists" = "CProp exists".
alias symbol "compose" = "bishop set relations composition".
for @{'cauchy $a}.
notation > "a 'is_cauchy'" non associative with precedence 45
for @{'cauchy $a}.
-interpretation "Cauchy sequence" 'cauchy s = (cauchy _ s).
+interpretation "Cauchy sequence" 'cauchy s = (cauchy ? s).
(* Definition 2.15 *)
definition uniform_converge ≝
notation > "a 'uniform_converges' x" non associative with precedence 45
for @{'uniform_converge $a $x}.
interpretation "Uniform convergence" 'uniform_converge s u =
- (uniform_converge _ s u).
+ (uniform_converge ? s u).
(* Lemma 2.16 *)
lemma uniform_converge_is_cauchy :
coercion Leibniz.
*)
-interpretation "setoid1 eq" 'eq t x y = (eq_rel1 _ (eq1 t) x y).
-interpretation "setoid eq" 'eq t x y = (eq_rel _ (eq t) x y).
-interpretation "setoid1 symmetry" 'invert r = (sym1 ____ r).
-interpretation "setoid symmetry" 'invert r = (sym ____ r).
+interpretation "setoid1 eq" 'eq t x y = (eq_rel1 ? (eq1 t) x y).
+interpretation "setoid eq" 'eq t x y = (eq_rel ? (eq t) x y).
+interpretation "setoid1 symmetry" 'invert r = (sym1 ???? r).
+interpretation "setoid symmetry" 'invert r = (sym ???? r).
notation ".= r" with precedence 50 for @{'trans $r}.
-interpretation "trans1" 'trans r = (trans1 _____ r).
-interpretation "trans" 'trans r = (trans _____ r).
+interpretation "trans1" 'trans r = (trans1 ????? r).
+interpretation "trans" 'trans r = (trans ????? r).
record unary_morphism (A,B: setoid1) : Type0 ≝
{ fun_1:1> A → B;
notation "† c" with precedence 90 for @{'prop1 $c }.
notation "l ‡ r" with precedence 90 for @{'prop $l $r }.
notation "#" with precedence 90 for @{'refl}.
-interpretation "prop_1" 'prop1 c = (prop_1 _____ c).
-interpretation "prop1" 'prop l r = (prop1 ________ l r).
-interpretation "prop" 'prop l r = (prop ________ l r).
-interpretation "refl1" 'refl = (refl1 ___).
-interpretation "refl" 'refl = (refl ___).
+interpretation "prop_1" 'prop1 c = (prop_1 ????? c).
+interpretation "prop1" 'prop l r = (prop1 ???????? l r).
+interpretation "prop" 'prop l r = (prop ???????? l r).
+interpretation "refl1" 'refl = (refl1 ???).
+interpretation "refl" 'refl = (refl ???).
definition CPROP: setoid1.
constructor 1;
qed.
notation ". r" with precedence 50 for @{'if $r}.
-interpretation "if" 'if r = (if' __ r).
+interpretation "if" 'if r = (if' ?? r).
definition and_morphism: binary_morphism1 CPROP CPROP CPROP.
constructor 1;
| apply (fi ?? H1 b1)]]
qed.
-interpretation "and_morphism" 'and a b = (fun1 ___ and_morphism a b).
+interpretation "and_morphism" 'and a b = (fun1 ??? and_morphism a b).
definition or_morphism: binary_morphism1 CPROP CPROP CPROP.
constructor 1;
| apply (fi ?? H1 b1)]]
qed.
-interpretation "or_morphism" 'or a b = (fun1 ___ or_morphism a b).
+interpretation "or_morphism" 'or a b = (fun1 ??? or_morphism a b).
definition if_morphism: binary_morphism1 CPROP CPROP CPROP.
constructor 1;
notation "'ASSOC'" with precedence 90 for @{'assoc}.
notation "'ASSOC1'" with precedence 90 for @{'assoc1}.
-interpretation "category1 composition" 'compose x y = (fun1 ___ (comp1 ____) y x).
-interpretation "category1 assoc" 'assoc1 = (comp_assoc1 ________).
-interpretation "category composition" 'compose x y = (fun ___ (comp ____) y x).
-interpretation "category assoc" 'assoc = (comp_assoc ________).
+interpretation "category1 composition" 'compose x y = (fun1 ??? (comp1 ????) y x).
+interpretation "category1 assoc" 'assoc1 = (comp_assoc1 ????????).
+interpretation "category composition" 'compose x y = (fun ??? (comp ????) y x).
+interpretation "category assoc" 'assoc = (comp_assoc ????????).
definition unary_morphism_setoid: setoid → setoid → setoid.
intros;
intros; apply (prop_1 A B w a b H);
qed.
-interpretation "SET dagger" 'prop1 h = (prop_1_SET _ _ _ _ _ h).
+interpretation "SET dagger" 'prop1 h = (prop_1_SET ? ? ? ? ? h).
inductive Prod (A,B:Type) : Type \def
pair : A \to B \to Prod A B.
-interpretation "Pair construction" 'pair x y = (pair _ _ x y).
+interpretation "Pair construction" 'pair x y = (pair ? ? x y).
interpretation "Product" 'product x y = (Prod x y).
match p with
[(pair a b) \Rightarrow b].
-interpretation "pair pi1" 'pi1 = (fst _ _).
-interpretation "pair pi2" 'pi2 = (snd _ _).
-interpretation "pair pi1" 'pi1a x = (fst _ _ x).
-interpretation "pair pi2" 'pi2a x = (snd _ _ x).
-interpretation "pair pi1" 'pi1b x y = (fst _ _ x y).
-interpretation "pair pi2" 'pi2b x y = (snd _ _ x y).
+interpretation "pair pi1" 'pi1 = (fst ? ?).
+interpretation "pair pi2" 'pi2 = (snd ? ?).
+interpretation "pair pi1" 'pi1a x = (fst ? ? x).
+interpretation "pair pi2" 'pi2a x = (snd ? ? x).
+interpretation "pair pi1" 'pi1b x y = (fst ? ? x y).
+interpretation "pair pi2" 'pi2b x y = (snd ? ? x y).
theorem eq_pair_fst_snd: \forall A,B:Type.\forall p:Prod A B.
p = 〈 \fst p, \snd p 〉.
| apply s1; assumption]]
qed.
-interpretation "mem" 'mem a S = (fun1 ___ (mem _) a S).
+interpretation "mem" 'mem a S = (fun1 ??? (mem ?) a S).
definition subseteq: ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) CPROP.
intros;
apply (transitive_subseteq_operator ???? s s4) ]]
qed.
-interpretation "subseteq" 'subseteq U V = (fun1 ___ (subseteq _) U V).
+interpretation "subseteq" 'subseteq U V = (fun1 ??? (subseteq ?) U V).
theorem subseteq_refl: ∀A.∀S:Ω \sup A.S ⊆ S.
intros 4; assumption.
| apply (. #‡(H1 \sup -1)); assumption]]
qed.
-interpretation "overlaps" 'overlaps U V = (fun1 ___ (overlaps _) U V).
+interpretation "overlaps" 'overlaps U V = (fun1 ??? (overlaps ?) U V).
definition intersects:
∀A. binary_morphism1 (powerset_setoid A) (powerset_setoid A) (powerset_setoid A).
| apply (. (#‡(H \sup -1))‡(#‡(H1 \sup -1))); assumption]]
qed.
-interpretation "intersects" 'intersects U V = (fun1 ___ (intersects _) U V).
+interpretation "intersects" 'intersects U V = (fun1 ??? (intersects ?) U V).
definition union:
∀A. binary_morphism1 (powerset_setoid A) (powerset_setoid A) (powerset_setoid A).
| apply (. (#‡(H \sup -1))‡(#‡(H1 \sup -1))); assumption]]
qed.
-interpretation "union" 'union U V = (fun1 ___ (union _) U V).
+interpretation "union" 'union U V = (fun1 ??? (union ?) U V).
definition singleton: ∀A:setoid. unary_morphism A (Ω \sup A).
intros; constructor 1;
[ apply a |4: apply a'] try assumption; apply sym; assumption]
qed.
-interpretation "singleton" 'singl a = (fun_1 __ (singleton _) a).
+interpretation "singleton" 'singl a = (fun_1 ?? (singleton ?) a).
*)
\ No newline at end of file
right associative with precedence 80
for @{'sig $x $h}.
-interpretation "sub_eqType" 'sig x h = (mk_sigma _ _ x h).
+interpretation "sub_eqType" 'sig x h = (mk_sigma ? ? x h).
(* restricting an eqType gives an eqType *)
lemma sigma_eq_dec : ∀d:eqType.∀p.∀x,y:sigma d p.
record powerset (A : Type) : Type ≝ { char : A → CProp }.
-interpretation "char" 'subset p = (mk_powerset _ p).
+interpretation "char" 'subset p = (mk_powerset ? p).
interpretation "pwset" 'powerset a = (powerset a).
-interpretation "in" 'mem a X = (char _ X a).
+interpretation "in" 'mem a X = (char ? X a).
definition subseteq ≝ λA.λu,v:\Omega \sup A.∀x.x ∈ u → x ∈ v.
-interpretation "subseteq" 'subseteq u v = (subseteq _ u v).
+interpretation "subseteq" 'subseteq u v = (subseteq ? u v).
definition overlaps ≝ λA.λU,V : Ω \sup A. exT2 ? (λx.x ∈ U) (λx.x ∈ V).
-interpretation "overlaps" 'overlaps u v = (overlaps _ u v).
+interpretation "overlaps" 'overlaps u v = (overlaps ? u v).
definition intersect ≝ λA.λu,v:Ω\sup A.{ y | y ∈ u ∧ y ∈ v }.
-interpretation "intersect" 'intersects u v = (intersect _ u v).
+interpretation "intersect" 'intersects u v = (intersect ? u v).
record axiom_set : Type ≝ {
A:> Type;
notation "hvbox(a break ◃ b)" non associative with precedence 45
for @{ 'covers $a $b }. (* a \ltri b *)
-interpretation "coversl" 'covers A U = (for_all _ U A (covers _ U)).
-interpretation "covers" 'covers a U = (covers _ U a).
+interpretation "coversl" 'covers A U = (for_all ? U A (covers ? U)).
+interpretation "covers" 'covers a U = (covers ? U a).
definition covers_elim ≝
λA:axiom_set.λU: \Omega \sup A.λP:\Omega \sup A.
notation "hvbox(a break ⋉ b)" non associative with precedence 45
for @{ 'fish $a $b }. (* a \ltimes b *)
-interpretation "fishl" 'fish A U = (ex_such _ U A (fish _ U)).
-interpretation "fish" 'fish a U = (fish _ U a).
+interpretation "fishl" 'fish A U = (ex_such ? U A (fish ? U)).
+interpretation "fish" 'fish a U = (fish ? U a).
let corec fish_rec (A:axiom_set) (U: \Omega \sup A)
(P: Ω \sup A) (H1: P ⊆ U)
definition leq ≝ λA:axiom_set.λa,b:A. a ◃ {y|b=y}.
-interpretation "covered by one" 'leq a b = (leq _ a b).
+interpretation "covered by one" 'leq a b = (leq ? a b).
theorem leq_refl: ∀A:axiom_set.∀a:A. a ≤ a.
intros;
definition uparrow ≝ λA:axiom_set.λa:A.mk_powerset ? (λb:A. a ≤ b).
-interpretation "uparrow" 'uparrow a = (uparrow _ a).
+interpretation "uparrow" 'uparrow a = (uparrow ? a).
definition downarrow ≝ λA:axiom_set.λU:Ω \sup A.mk_powerset ? (λa:A. (↑a) ≬ U).
-interpretation "downarrow" 'downarrow a = (downarrow _ a).
+interpretation "downarrow" 'downarrow a = (downarrow ? a).
definition fintersects ≝ λA:axiom_set.λU,V:Ω \sup A.↓U ∩ ↓V.
-interpretation "fintersects" 'fintersects U V = (fintersects _ U V).
+interpretation "fintersects" 'fintersects U V = (fintersects ? U V).
record convergent_generated_topology : Type ≝
{ AA:> axiom_set;
apply rule (prove (A ∨ ¬ A));
apply rule (RAA [H] ⊥);
- apply rule (¬_e (¬A) A);
- [ apply rule (¬_i [H1] ⊥);
- apply rule (¬_e (¬(A∨¬A)) (A∨¬A));
+ apply rule (¬#e (¬A) A);
+ [ apply rule (¬#i [H1] ⊥);
+ apply rule (¬#e (¬(A∨¬A)) (A∨¬A));
[ apply rule (discharge [H]);
- | apply rule (∨_i_l A);
+ | apply rule (∨#i_l A);
apply rule (discharge [H1]);
]
| apply rule (RAA [H2] ⊥);
- apply rule (¬_e (¬(A∨¬A)) (A∨¬A));
+ apply rule (¬#e (¬(A∨¬A)) (A∨¬A));
[ apply rule (discharge [H]);
- | apply rule (∨_i_r (¬A));
+ | apply rule (∨#i_r (¬A));
apply rule (discharge [H2]);
]
]
apply rule (prove (A ∨ ¬ A));
apply rule (RAA [H] ⊥);
- apply rule (¬_e (¬¬A) (¬A));
- [ apply rule (¬_i [H2] ⊥);
- apply rule (¬_e (¬(A∨¬A)) (A∨¬A));
+ apply rule (¬#e (¬¬A) (¬A));
+ [ apply rule (¬#i [H2] ⊥);
+ apply rule (¬#e (¬(A∨¬A)) (A∨¬A));
[ apply rule (discharge [H]);
- | apply rule (∨_i_r (¬A));
+ | apply rule (∨#i_r (¬A));
apply rule (discharge [H2]);
]
- | apply rule (¬_i [H1] ⊥);
- apply rule (¬_e (¬(A∨¬A)) (A∨¬A));
+ | apply rule (¬#i [H1] ⊥);
+ apply rule (¬#e (¬(A∨¬A)) (A∨¬A));
[ apply rule (discharge [H]);
- | apply rule (∨_i_l A);
+ | apply rule (∨#i_l A);
apply rule (discharge [H1]);
]
]
apply rule (prove ((A⇒E)∨B⇒A∧C⇒E∧C∨B));
- apply rule (⇒_i [H] (A∧C⇒E∧C∨B));
- apply rule (⇒_i [K] (E∧C∨B));
- apply rule (∨_e ((A⇒E)∨B) [C1] (E∧C∨B) [C2] (E∧C∨B));
+ apply rule (⇒#i [H] (A∧C⇒E∧C∨B));
+ apply rule (⇒#i [K] (E∧C∨B));
+ apply rule (∨#e ((A⇒E)∨B) [C1] (E∧C∨B) [C2] (E∧C∨B));
[ apply rule (discharge [H]);
-| apply rule (∨_i_l (E∧C));
- apply rule (∧_i E C);
- [ apply rule (⇒_e (A⇒E) A);
+| apply rule (∨#i_l (E∧C));
+ apply rule (∧#i E C);
+ [ apply rule (⇒#e (A⇒E) A);
[ apply rule (discharge [C1]);
- | apply rule (∧_e_l (A∧C)); apply rule (discharge [K]);
+ | apply rule (∧#e_l (A∧C)); apply rule (discharge [K]);
]
- | apply rule (∧_e_r (A∧C)); apply rule (discharge [K]);
+ | apply rule (∧#e_r (A∧C)); apply rule (discharge [K]);
]
-| apply rule (∨_i_r B); apply rule (discharge [C2]);
+| apply rule (∨#i_r B); apply rule (discharge [C2]);
]
qed.
@{\lambda ${ident i} : $ty. $p}
@{\lambda ${ident i} . $p}}}.
-interpretation "Sigma" 'sigma \eta.x = (sigma _ x).
+interpretation "Sigma" 'sigma \eta.x = (sigma ? x).
let rec type_OF_SP F ≝
match F return λF.Type with
definition if_then_else ≝ λT:Type.λe,t,f.match e return λ_.T with [ true ⇒ t | false ⇒ f].
notation > "'if' term 19 e 'then' term 19 t 'else' term 90 f" non associative with precedence 19 for @{ 'if_then_else $e $t $f }.
notation < "'if' \nbsp term 19 e \nbsp 'then' \nbsp term 19 t \nbsp 'else' \nbsp term 90 f \nbsp" non associative with precedence 19 for @{ 'if_then_else $e $t $f }.
-interpretation "Formula if_then_else" 'if_then_else e t f = (if_then_else _ e t f).
+interpretation "Formula if_then_else" 'if_then_else e t f = (if_then_else ? e t f).
definition max ≝ λn,m. if eqb (n - m) 0 then m else n.
definition min ≝ λn,m. if eqb (n - m) 0 then n else m.
*)
notation < "[[ \nbsp term 19 a \nbsp ]] \nbsp \sub term 90 v" non associative with precedence 90 for @{ 'semantics $v $a }.
notation > "[[ term 19 a ]] \sub term 90 v" non associative with precedence 90 for @{ 'semantics $v $a }.
-notation > "[[ term 19 a ]]_ term 90 v" non associative with precedence 90 for @{ sem $v $a }.
+notation > "[[ term 19 a ]] term 90 v" non associative with precedence 90 for @{ sem $v $a }.
interpretation "Semantic of Formula" 'semantics v a = (sem v a).
definition v20 ≝ λx.
Decommenta ed esegui.
*)
-(* eval normalize on [[FOr (FAtom 0) (FAtom 2)]]_v20. *)
+(* eval normalize on [[FOr (FAtom 0) (FAtom 2)]]v20. *)
(*DOCBEGIN
librerie di teoremi già dimostrati. Per portare a termine l'esercitazione
sono necessari i seguenti lemmi:
-* lemma `sem_le_1` : `∀F,v. [[ F ]]_v ≤ 1`
+* lemma `sem_le_1` : `∀F,v. [[ F ]]v ≤ 1`
* lemma `min_1_1` : `∀x. x ≤ 1 → 1 - (1 - x) = x`
* lemma `min_bool` : `∀n. min n 1 = 0 ∨ min n 1 = 1`
-* lemma `min_max` : `∀F,G,v.min (1 - [[F]]_v) (1 - [[G]]_v) = 1 - max [[F]]_v [[G]]_v`
-* lemma `max_min` : `∀F,G,v.max (1 - [[F]]_v) (1 - [[G]]_v) = 1 - min [[F]]_v [[G]]_v`
+* lemma `min_max` : `∀F,G,v.min (1 - [[F]]v) (1 - [[G]]v) = 1 - max [[F]]v [[G]]v`
+* lemma `max_min` : `∀F,G,v.max (1 - [[F]]v) (1 - [[G]]v) = 1 - min [[F]]v [[G]]v`
* lemma `equiv_rewrite` : `∀F1,F2,F3. F1 ≡ F2 → F1 ≡ F3 → F2 ≡ F3`
* lemma `equiv_sym` : `∀F1,F2. F1 ≡ F2 → F2 ≡ F1`
Non modificare quanto segue.
*)
-lemma sem_bool : ∀F,v. [[ F ]]_v = 0 ∨ [[ F ]]_v = 1. intros; elim F; simplify; [left;reflexivity; |right;reflexivity; |cases (v n);[left;|cases n1;right;]reflexivity; |4,5,6: cases H; cases H1; rewrite > H2; rewrite > H3; simplify; first [ left;reflexivity | right; reflexivity ]. |cases H; rewrite > H1; simplify;[right|left]reflexivity;] qed.
+lemma sem_bool : ∀F,v. [[ F ]]v = 0 ∨ [[ F ]]v = 1. intros; elim F; simplify; [left;reflexivity; |right;reflexivity; |cases (v n);[left;|cases n1;right;]reflexivity; |4,5,6: cases H; cases H1; rewrite > H2; rewrite > H3; simplify; first [ left;reflexivity | right; reflexivity ]. |cases H; rewrite > H1; simplify;[right|left]reflexivity;] qed.
lemma min_bool : ∀n. min n 1 = 0 ∨ min n 1 = 1. intros; cases n; [left;reflexivity] cases n1; right; reflexivity; qed.
-lemma min_max : ∀F,G,v. min (1 - [[F]]_v) (1 - [[G]]_v) = 1 - max [[F]]_v [[G]]_v. intros; cases (sem_bool F v);cases (sem_bool G v); rewrite > H; rewrite >H1; simplify; reflexivity; qed.
-lemma max_min : ∀F,G,v. max (1 - [[F]]_v) (1 - [[G]]_v) = 1 - min [[F]]_v [[G]]_v. intros; cases (sem_bool F v);cases (sem_bool G v); rewrite > H; rewrite >H1; simplify; reflexivity; qed.
+lemma min_max : ∀F,G,v. min (1 - [[F]]v) (1 - [[G]]v) = 1 - max [[F]]v [[G]]v. intros; cases (sem_bool F v);cases (sem_bool G v); rewrite > H; rewrite >H1; simplify; reflexivity; qed.
+lemma max_min : ∀F,G,v. max (1 - [[F]]v) (1 - [[G]]v) = 1 - min [[F]]v [[G]]v. intros; cases (sem_bool F v);cases (sem_bool G v); rewrite > H; rewrite >H1; simplify; reflexivity; qed.
lemma min_1_1 : ∀x.x ≤ 1 → 1 - (1 - x) = x. intros; inversion H; intros; destruct; [reflexivity;] rewrite < (le_n_O_to_eq ? H1); reflexivity;qed.
-lemma sem_le_1 : ∀F,v.[[F]]_v ≤ 1. intros; cases (sem_bool F v); rewrite > H; [apply le_O_n|apply le_n]qed.
+lemma sem_le_1 : ∀F,v.[[F]]v ≤ 1. intros; cases (sem_bool F v); rewrite > H; [apply le_O_n|apply le_n]qed.
(* Esercizio 2
===========
Non modificare quanto segue
*)
-definition equiv ≝ λF1,F2. ∀v.[[ F1 ]]_v = [[ F2 ]]_v.
+definition equiv ≝ λF1,F2. ∀v.[[ F1 ]]v = [[ F2 ]]v.
notation "hvbox(a \nbsp break mstyle color #0000ff (≡) \nbsp b)" non associative with precedence 45 for @{ 'equivF $a $b }.
notation > "a ≡ b" non associative with precedence 50 for @{ equiv $a $b }.
interpretation "equivalence for Formulas" 'equivF a b = (equiv a b).
a `F` in un mondo invertito.
*)
lemma negate_invert:
- ∀F:Formula.∀v:ℕ→ℕ.[[ negate F ]]_v=[[ F ]]_(invert v).
+ ∀F:Formula.∀v:ℕ→ℕ.[[ negate F ]]v=[[ F ]](invert v).
assume F:Formula.
assume v:(ℕ→ℕ).
-we proceed by induction on F to prove ([[ negate F ]]_v=[[ F ]]_(invert v)).
+we proceed by induction on F to prove ([[ negate F ]]v=[[ F ]](invert v)).
case FBot.
(*BEGIN*)
- the thesis becomes ([[ negate FBot ]]_v=[[ FBot ]]_(invert v)).
+ the thesis becomes ([[ negate FBot ]]v=[[ FBot ]](invert v)).
(*END*)
done.
case FTop.
(*BEGIN*)
- the thesis becomes ([[ negate FTop ]]_v=[[ FTop ]]_(invert v)).
+ the thesis becomes ([[ negate FTop ]]v=[[ FTop ]](invert v)).
(*END*)
done.
case FAtom.
assume n : ℕ.
- the thesis becomes ((*BEGIN*)[[ negate (FAtom n) ]]_v=[[ FAtom n ]]_(invert v)(*END*)).
+ the thesis becomes ((*BEGIN*)[[ negate (FAtom n) ]]v=[[ FAtom n ]](invert v)(*END*)).
the thesis becomes ((*BEGIN*)1 - (min (v n) 1)= min (invert v n) 1(*END*)).
the thesis becomes (1 - (min (v n) 1)= min (if eqb (min (v n) 1) 0 then 1 else 0) 1).
by min_bool we proved (min (v n) 1 = 0 ∨ min (v n) 1 = 1) (H1);
case FAnd.
assume f : Formula.
by induction hypothesis we know
- ((*BEGIN*)[[ negate f ]]_v=[[ f ]]_(invert v)(*END*)) (H).
+ ((*BEGIN*)[[ negate f ]]v=[[ f ]](invert v)(*END*)) (H).
assume f1 : Formula.
by induction hypothesis we know
- ((*BEGIN*)[[ negate f1 ]]_v=[[ f1 ]]_(invert v)(*END*)) (H1).
+ ((*BEGIN*)[[ negate f1 ]]v=[[ f1 ]](invert v)(*END*)) (H1).
the thesis becomes
- ([[ negate (FAnd f f1) ]]_v=[[ FAnd f f1 ]]_(invert v)).
+ ([[ negate (FAnd f f1) ]]v=[[ FAnd f f1 ]](invert v)).
the thesis becomes
- (min [[ negate f ]]_v [[ negate f1]]_v = [[ FAnd f f1 ]]_(invert v)).
+ (min [[ negate f ]]v [[ negate f1]]v = [[ FAnd f f1 ]](invert v)).
conclude
- (min [[ negate f ]]_v [[ negate f1]]_v)
- = (min [[ f ]]_(invert v) [[ negate f1]]_v) by (*BEGIN*)H(*END*).
- = (min [[ f ]]_(invert v) [[ f1]]_(invert v)) by (*BEGIN*)H1(*END*).
+ (min [[ negate f ]]v [[ negate f1]]v)
+ = (min [[ f ]](invert v) [[ negate f1]]v) by (*BEGIN*)H(*END*).
+ = (min [[ f ]](invert v) [[ f1]](invert v)) by (*BEGIN*)H1(*END*).
done.
case FOr.
(*BEGIN*)
assume f : Formula.
by induction hypothesis we know
- ([[ negate f ]]_v=[[ f ]]_(invert v)) (H).
+ ([[ negate f ]]v=[[ f ]](invert v)) (H).
assume f1 : Formula.
by induction hypothesis we know
- ([[ negate f1 ]]_v=[[ f1 ]]_(invert v)) (H1).
+ ([[ negate f1 ]]v=[[ f1 ]](invert v)) (H1).
the thesis becomes
- ([[ negate (FOr f f1) ]]_v=[[ FOr f f1 ]]_(invert v)).
+ ([[ negate (FOr f f1) ]]v=[[ FOr f f1 ]](invert v)).
the thesis becomes
- (max [[ negate f ]]_v [[ negate f1]]_v = [[ FOr f f1 ]]_(invert v)).
+ (max [[ negate f ]]v [[ negate f1]]v = [[ FOr f f1 ]](invert v)).
conclude
- (max [[ negate f ]]_v [[ negate f1]]_v)
- = (max [[ f ]]_(invert v) [[ negate f1]]_v) by H.
- = (max [[ f ]]_(invert v) [[ f1]]_(invert v)) by H1.
+ (max [[ negate f ]]v [[ negate f1]]v)
+ = (max [[ f ]](invert v) [[ negate f1]]v) by H.
+ = (max [[ f ]](invert v) [[ f1]](invert v)) by H1.
(*END*)
done.
case FImpl.
(*BEGIN*)
assume f : Formula.
by induction hypothesis we know
- ([[ negate f ]]_v=[[ f ]]_(invert v)) (H).
+ ([[ negate f ]]v=[[ f ]](invert v)) (H).
assume f1 : Formula.
by induction hypothesis we know
- ([[ negate f1 ]]_v=[[ f1 ]]_(invert v)) (H1).
+ ([[ negate f1 ]]v=[[ f1 ]](invert v)) (H1).
the thesis becomes
- ([[ negate (FImpl f f1) ]]_v=[[ FImpl f f1 ]]_(invert v)).
+ ([[ negate (FImpl f f1) ]]v=[[ FImpl f f1 ]](invert v)).
the thesis becomes
- (max (1 - [[ negate f ]]_v) [[ negate f1]]_v = [[ FImpl f f1 ]]_(invert v)).
+ (max (1 - [[ negate f ]]v) [[ negate f1]]v = [[ FImpl f f1 ]](invert v)).
conclude
- (max (1 - [[ negate f ]]_v) [[ negate f1]]_v)
- = (max (1 - [[ f ]]_(invert v)) [[ negate f1]]_v) by H.
- = (max (1 - [[ f ]]_(invert v)) [[ f1]]_(invert v)) by H1.
+ (max (1 - [[ negate f ]]v) [[ negate f1]]v)
+ = (max (1 - [[ f ]](invert v)) [[ negate f1]]v) by H.
+ = (max (1 - [[ f ]](invert v)) [[ f1]](invert v)) by H1.
(*END*)
done.
case FNot.
(*BEGIN*)
assume f : Formula.
by induction hypothesis we know
- ([[ negate f ]]_v=[[ f ]]_(invert v)) (H).
+ ([[ negate f ]]v=[[ f ]](invert v)) (H).
the thesis becomes
- ([[ negate (FNot f) ]]_v=[[ FNot f ]]_(invert v)).
+ ([[ negate (FNot f) ]]v=[[ FNot f ]](invert v)).
the thesis becomes
- (1 - [[ negate f ]]_v=[[ FNot f ]]_(invert v)).
- conclude (1 - [[ negate f ]]_v) = (1 - [[f]]_(invert v)) by H.
+ (1 - [[ negate f ]]v=[[ FNot f ]](invert v)).
+ conclude (1 - [[ negate f ]]v) = (1 - [[f]](invert v)) by H.
(*END*)
done.
qed.
assume (*BEGIN*)G:Formula(*END*).
suppose (*BEGIN*)(F ≡ G) (H)(*END*).
the thesis becomes (*BEGIN*)(negate F ≡ negate G)(*END*).
- the thesis becomes (*BEGIN*)(∀v:ℕ→ℕ.[[ negate F ]]_v=[[ negate G ]]_v)(*END*).
+ the thesis becomes (*BEGIN*)(∀v:ℕ→ℕ.[[ negate F ]]v=[[ negate G ]]v)(*END*).
assume v:(ℕ→ℕ).
conclude
- [[ negate F ]]_v
- = [[ F ]]_(invert v) by (*BEGIN*)negate_invert(*END*).
- = [[ G ]]_((*BEGIN*)invert v(*BEGIN*)) by (*BEGIN*)H(*BEGIN*).
- = [[ negate G ]]_(*BEGIN*)v(*BEGIN*) by (*BEGIN*)negate_invert(*END*).
+ [[ negate F ]]v
+ = [[ F ]](invert v) by (*BEGIN*)negate_invert(*END*).
+ = [[ G ]]((*BEGIN*)invert v(*BEGIN*)) by (*BEGIN*)H(*BEGIN*).
+ = [[ negate G ]](*BEGIN*)v(*BEGIN*) by (*BEGIN*)negate_invert(*END*).
done.
qed.
∀F:Formula.negate F ≡ FNot (dualize F).
(*BEGIN*)
assume F:Formula.
- the thesis becomes (∀v:ℕ→ℕ.[[negate F]]_v=[[FNot (dualize F)]]_v).
+ the thesis becomes (∀v:ℕ→ℕ.[[negate F]]v=[[FNot (dualize F)]]v).
(*END*)
assume v:(ℕ→ℕ).
- we proceed by induction on F to prove ([[negate F]]_v=[[FNot (dualize F)]]_v).
+ we proceed by induction on F to prove ([[negate F]]v=[[FNot (dualize F)]]v).
case FBot.
(*BEGIN*)
- the thesis becomes ([[ negate FBot ]]_v=[[ FNot (dualize FBot) ]]_v).
+ the thesis becomes ([[ negate FBot ]]v=[[ FNot (dualize FBot) ]]v).
(*END*)
done.
case FTop.
(*BEGIN*)
- the thesis becomes ([[ negate FTop ]]_v=[[ FNot (dualize FTop) ]]_v).
+ the thesis becomes ([[ negate FTop ]]v=[[ FNot (dualize FTop) ]]v).
(*END*)
done.
case FAtom.
(*BEGIN*)
assume n : ℕ.
- the thesis becomes ([[ negate (FAtom n) ]]_v=[[ FNot (dualize (FAtom n)) ]]_v).
+ the thesis becomes ([[ negate (FAtom n) ]]v=[[ FNot (dualize (FAtom n)) ]]v).
(*END*)
done.
case FAnd.
assume f : Formula.
by induction hypothesis we know
- ([[ negate f ]]_v=[[ FNot (dualize f) ]]_v) (H).
+ ([[ negate f ]]v=[[ FNot (dualize f) ]]v) (H).
assume f1 : Formula.
by induction hypothesis we know
- ([[ negate f1 ]]_v=[[ FNot (dualize f1) ]]_v) (H1).
+ ([[ negate f1 ]]v=[[ FNot (dualize f1) ]]v) (H1).
the thesis becomes
- ([[ negate (FAnd f f1) ]]_v=[[ FNot (dualize (FAnd f f1)) ]]_v).
+ ([[ negate (FAnd f f1) ]]v=[[ FNot (dualize (FAnd f f1)) ]]v).
the thesis becomes
- (min [[ negate f ]]_v [[ negate f1 ]]_v=[[ FNot (dualize (FAnd f f1)) ]]_v).
+ (min [[ negate f ]]v [[ negate f1 ]]v=[[ FNot (dualize (FAnd f f1)) ]]v).
conclude
- (min (*BEGIN*)[[ negate f ]]_v(*END*) (*BEGIN*)[[ negate f1 ]]_v(*END*))
- = (min (*BEGIN*)[[ FNot (dualize f) ]]_v(*END*) (*BEGIN*)[[ negate f1 ]]_v(*END*)) by H.
- = (min (*BEGIN*)[[ FNot (dualize f) ]]_v(*END*) (*BEGIN*)[[ FNot (dualize f1) ]]_v(*END*)) by H1.
- = (min (1 - [[ dualize f ]]_v) (1 - [[ dualize f1 ]]_v)).
- = (1 - (max [[ dualize f ]]_v [[ dualize f1 ]]_v)) by min_max.
+ (min (*BEGIN*)[[ negate f ]]v(*END*) (*BEGIN*)[[ negate f1 ]]v(*END*))
+ = (min (*BEGIN*)[[ FNot (dualize f) ]]v(*END*) (*BEGIN*)[[ negate f1 ]]v(*END*)) by H.
+ = (min (*BEGIN*)[[ FNot (dualize f) ]]v(*END*) (*BEGIN*)[[ FNot (dualize f1) ]]v(*END*)) by H1.
+ = (min (1 - [[ dualize f ]]v) (1 - [[ dualize f1 ]]v)).
+ = (1 - (max [[ dualize f ]]v [[ dualize f1 ]]v)) by min_max.
done.
case FOr.
(*BEGIN*)
assume f : Formula.
by induction hypothesis we know
- ([[ negate f ]]_v=[[ FNot (dualize f) ]]_v) (H).
+ ([[ negate f ]]v=[[ FNot (dualize f) ]]v) (H).
assume f1 : Formula.
by induction hypothesis we know
- ([[ negate f1 ]]_v=[[ FNot (dualize f1) ]]_v) (H1).
+ ([[ negate f1 ]]v=[[ FNot (dualize f1) ]]v) (H1).
the thesis becomes
- ([[ negate (FOr f f1) ]]_v=[[ FNot (dualize (FOr f f1)) ]]_v).
+ ([[ negate (FOr f f1) ]]v=[[ FNot (dualize (FOr f f1)) ]]v).
the thesis becomes
- (max [[ negate f ]]_v [[ negate f1 ]]_v=[[ FNot (dualize (FOr f f1)) ]]_v).
+ (max [[ negate f ]]v [[ negate f1 ]]v=[[ FNot (dualize (FOr f f1)) ]]v).
conclude
- (max [[ negate f ]]_v [[ negate f1 ]]_v)
- = (max [[ FNot (dualize f) ]]_v [[ negate f1 ]]_v) by H.
- = (max [[ FNot (dualize f) ]]_v [[ FNot (dualize f1) ]]_v) by H1.
- = (max (1 - [[ dualize f ]]_v) (1 - [[ dualize f1 ]]_v)).
- = (1 - (min [[ dualize f ]]_v [[ dualize f1 ]]_v)) by max_min.
+ (max [[ negate f ]]v [[ negate f1 ]]v)
+ = (max [[ FNot (dualize f) ]]v [[ negate f1 ]]v) by H.
+ = (max [[ FNot (dualize f) ]]v [[ FNot (dualize f1) ]]v) by H1.
+ = (max (1 - [[ dualize f ]]v) (1 - [[ dualize f1 ]]v)).
+ = (1 - (min [[ dualize f ]]v [[ dualize f1 ]]v)) by max_min.
(*END*)
done.
case FImpl.
(*BEGIN*)
assume f : Formula.
by induction hypothesis we know
- ([[ negate f ]]_v=[[ FNot (dualize f) ]]_v) (H).
+ ([[ negate f ]]v=[[ FNot (dualize f) ]]v) (H).
assume f1 : Formula.
by induction hypothesis we know
- ([[ negate f1 ]]_v=[[ FNot (dualize f1) ]]_v) (H1).
+ ([[ negate f1 ]]v=[[ FNot (dualize f1) ]]v) (H1).
the thesis becomes
- ([[ negate (FImpl f f1) ]]_v=[[ FNot (dualize (FImpl f f1)) ]]_v).
+ ([[ negate (FImpl f f1) ]]v=[[ FNot (dualize (FImpl f f1)) ]]v).
the thesis becomes
- (max (1 - [[ negate f ]]_v) [[ negate f1 ]]_v=[[ FNot (dualize (FImpl f f1)) ]]_v).
+ (max (1 - [[ negate f ]]v) [[ negate f1 ]]v=[[ FNot (dualize (FImpl f f1)) ]]v).
conclude
- (max (1-[[ negate f ]]_v) [[ negate f1 ]]_v)
- = (max (1-[[ FNot (dualize f) ]]_v) [[ negate f1 ]]_v) by H.
- = (max (1-[[ FNot (dualize f) ]]_v) [[ FNot (dualize f1) ]]_v) by H1.
- = (max (1 - [[ FNot (dualize f) ]]_v) (1 - [[ dualize f1 ]]_v)).
- = (1 - (min [[ FNot (dualize f) ]]_v [[ dualize f1 ]]_v)) by max_min.
+ (max (1-[[ negate f ]]v) [[ negate f1 ]]v)
+ = (max (1-[[ FNot (dualize f) ]]v) [[ negate f1 ]]v) by H.
+ = (max (1-[[ FNot (dualize f) ]]v) [[ FNot (dualize f1) ]]v) by H1.
+ = (max (1 - [[ FNot (dualize f) ]]v) (1 - [[ dualize f1 ]]v)).
+ = (1 - (min [[ FNot (dualize f) ]]v [[ dualize f1 ]]v)) by max_min.
(*END*)
done.
case FNot.
(*BEGIN*)
assume f : Formula.
by induction hypothesis we know
- ([[ negate f ]]_v=[[ FNot (dualize f) ]]_v) (H).
+ ([[ negate f ]]v=[[ FNot (dualize f) ]]v) (H).
the thesis becomes
- ([[ negate (FNot f) ]]_v=[[ FNot (dualize (FNot f)) ]]_v).
+ ([[ negate (FNot f) ]]v=[[ FNot (dualize (FNot f)) ]]v).
the thesis becomes
- (1 - [[ negate f ]]_v=[[ FNot (dualize (FNot f)) ]]_v).
- conclude (1 - [[ negate f ]]_v) = (1 - [[ FNot (dualize f) ]]_v) by H.
+ (1 - [[ negate f ]]v=[[ FNot (dualize (FNot f)) ]]v).
+ conclude (1 - [[ negate f ]]v) = (1 - [[ FNot (dualize f) ]]v) by H.
(*END*)
done.
qed.
assume G:Formula.
suppose (FNot F ≡ FNot G) (H).
the thesis becomes (F ≡ G).
- the thesis becomes (∀v:ℕ→ℕ.[[ F ]]_v=[[ G ]]_v).
+ the thesis becomes (∀v:ℕ→ℕ.[[ F ]]v=[[ G ]]v).
(*END*)
assume v:(ℕ→ℕ).
- by sem_le_1 we proved ([[F]]_v ≤ 1) (H1).
- by (*BEGIN*)sem_le_1(*END*) we proved ([[G]]_v ≤ 1) (H2).
- by min_1_1, H1 we proved (1 - (1 - [[F]]_v) = [[F]]_v) (H3).
- by (*BEGIN*)min_1_1, H2(*END*) we proved ((*BEGIN*)1 - (1 - [[G]]_v)(*END*) = [[G]]_v) (H4).
+ by sem_le_1 we proved ([[F]]v ≤ 1) (H1).
+ by (*BEGIN*)sem_le_1(*END*) we proved ([[G]]v ≤ 1) (H2).
+ by min_1_1, H1 we proved (1 - (1 - [[F]]v) = [[F]]v) (H3).
+ by (*BEGIN*)min_1_1, H2(*END*) we proved ((*BEGIN*)1 - (1 - [[G]]v)(*END*) = [[G]]v) (H4).
conclude
- ([[F]]_v)
- = (1 - (1 - [[F]]_v)) by (*BEGIN*)H3(*END*).
- = (1 - [[(*BEGIN*)FNot F(*END*)]]_v).
- = (1 - [[ FNot G]]_v) by H.
- = (1 - (*BEGIN*)(1 - [[G]]_v)(*END*)).
- = [[G]]_v by (*BEGIN*)H4(*END*).
+ ([[F]]v)
+ = (1 - (1 - [[F]]v)) by (*BEGIN*)H3(*END*).
+ = (1 - [[(*BEGIN*)FNot F(*END*)]]v).
+ = (1 - [[ FNot G]]v) by H.
+ = (1 - (*BEGIN*)(1 - [[G]]v)(*END*)).
+ = [[G]]v by (*BEGIN*)H4(*END*).
done.
qed.
1. lemma `negate_invert`, dimostrato per induzione su F, utilizzando
`min_bool`
- ∀F:Formula.∀v:ℕ→ℕ.[[ negate F ]]_v=[[ F ]]_(invert v).
+ ∀F:Formula.∀v:ℕ→ℕ.[[ negate F ]]v=[[ F ]]_(invert v).
2. lemma `negate_fun`, conseguenza di `negate_invert`
con premesse multiple sono seguite da `[`, `|` e `]`.
Ad esempio:
- apply rule (∧_i (A∨B) ⊥);
+ apply rule (∧#i (A∨B) ⊥);
[ …continua qui il sottoalbero per (A∨B)
| …continua qui il sottoalbero per ⊥
]
scrivete sopra la linea che rappresenta l'applicazione della
regola stessa. Ad esempio:
- aply rule (∨_e (A∨B) [h1] C [h2] C);
+ aply rule (∨#e (A∨B) [h1] C [h2] C);
[ …prova di (A∨B)
| …prova di C sotto l'ipotesi A (di nome h1)
| …prova di C sotto l'ipotesi B (di nome h2)
(* qui inizia l'albero, eseguite passo passo osservando come
si modifica l'albero. *)
apply rule (RAA [H] (⊥)).
-apply rule (¬_e (¬(A ∨ ¬A)) (A ∨ ¬A));
+apply rule (¬#e (¬(A ∨ ¬A)) (A ∨ ¬A));
[ apply rule (discharge [H]).
- | apply rule (⊥_e (⊥));
- apply rule (¬_e (¬¬A) (¬A));
- [ apply rule (¬_i [K] (⊥)).
- apply rule (¬_e (¬(A ∨ ¬A)) (A ∨ ¬A));
+ | apply rule (⊥#e (⊥));
+ apply rule (¬#e (¬¬A) (¬A));
+ [ apply rule (¬#i [K] (⊥)).
+ apply rule (¬#e (¬(A ∨ ¬A)) (A ∨ ¬A));
[ (*BEGIN*)apply rule (discharge [H]).(*END*)
- | (*BEGIN*)apply rule (∨_i_r (¬A)).
+ | (*BEGIN*)apply rule (∨#i_r (¬A)).
apply rule (discharge [K]).(*END*)
]
- | (*BEGIN*)apply rule (¬_i [K] (⊥)).
- apply rule (¬_e (¬(A ∨ ¬A)) (A ∨ ¬A));
+ | (*BEGIN*)apply rule (¬#i [K] (⊥)).
+ apply rule (¬#e (¬(A ∨ ¬A)) (A ∨ ¬A));
[ apply rule (discharge [H]).
- | apply rule (∨_i_l (A)).
+ | apply rule (∨#i_l (A)).
apply rule (discharge [K]).
](*END*)
]
theorem ex1 : (C∧G ⇒ E) ⇒ (¬L ⇒ E∨C) ⇒ G ∨ L ⇒ ¬L ⇒ E.
apply rule (prove ((C∧G ⇒ E) ⇒ (¬L ⇒ E∨C) ⇒ G ∨ L ⇒ ¬L ⇒ E));
(*BEGIN*)
-apply rule (⇒_i [h1] ((¬L ⇒ E∨C) ⇒ G ∨ L ⇒ ¬L ⇒ E));
-apply rule (⇒_i [h2] (G ∨ L ⇒ ¬L ⇒ E));
-apply rule (⇒_i [h3] (¬L ⇒ E));
-apply rule (⇒_i [h4] (E));
-apply rule (∨_e (G∨L) [h5] (E) [h6] (E));
+apply rule (⇒#i [h1] ((¬L ⇒ E∨C) ⇒ G ∨ L ⇒ ¬L ⇒ E));
+apply rule (⇒#i [h2] (G ∨ L ⇒ ¬L ⇒ E));
+apply rule (⇒#i [h3] (¬L ⇒ E));
+apply rule (⇒#i [h4] (E));
+apply rule (∨#e (G∨L) [h5] (E) [h6] (E));
[ apply rule (discharge [h3]);
- | apply rule (∨_e (E∨C) [h6] (E) [h7] (E));
- [ apply rule (⇒_e (¬L ⇒ E∨C) (¬L));
+ | apply rule (∨#e (E∨C) [h6] (E) [h7] (E));
+ [ apply rule (⇒#e (¬L ⇒ E∨C) (¬L));
[ apply rule (discharge [h2]);
| apply rule (discharge [h4]);
]
| apply rule (discharge [h6]);
- | apply rule (⇒_e (C∧G ⇒ E) (C∧G));
+ | apply rule (⇒#e (C∧G ⇒ E) (C∧G));
[ apply rule (discharge [h1]);
- | apply rule (∧_i (C) (G));
+ | apply rule (∧#i (C) (G));
[ apply rule (discharge [h7]);
| apply rule (discharge [h5]);
]
]
]
- | apply rule (⊥_e (⊥));
- apply rule (¬_e (¬L) (L));
+ | apply rule (⊥#e (⊥));
+ apply rule (¬#e (¬L) (L));
[ apply rule (discharge [h4]);
| apply rule (discharge [h6]);
]
theorem ex2 : (A∧¬B ⇒ C) ⇒ (B∧D ⇒ C) ⇒ (D ⇒ A) ⇒ D ⇒ C.
apply rule (prove ((A∧¬B ⇒ C) ⇒ (B∧D ⇒ C) ⇒ (D ⇒ A) ⇒ D ⇒ C));
(*BEGIN*)
-apply rule (⇒_i [h1] ((B∧D ⇒ C) ⇒ (D ⇒ A) ⇒ D ⇒ C));
-apply rule (⇒_i [h2] ((D ⇒ A) ⇒ D ⇒ C));
-apply rule (⇒_i [h3] (D ⇒ C));
-apply rule (⇒_i [h4] (C));
-apply rule (∨_e (B∨¬B) [h5] (C) [h6] (C));
+apply rule (⇒#i [h1] ((B∧D ⇒ C) ⇒ (D ⇒ A) ⇒ D ⇒ C));
+apply rule (⇒#i [h2] ((D ⇒ A) ⇒ D ⇒ C));
+apply rule (⇒#i [h3] (D ⇒ C));
+apply rule (⇒#i [h4] (C));
+apply rule (∨#e (B∨¬B) [h5] (C) [h6] (C));
[ apply rule (lem 0 EM);
- | apply rule (⇒_e (B∧D⇒C) (B∧D));
+ | apply rule (⇒#e (B∧D⇒C) (B∧D));
[ apply rule (discharge [h2]);
- | apply rule (∧_i (B) (D));
+ | apply rule (∧#i (B) (D));
[ apply rule (discharge [h5]);
| apply rule (discharge [h4]);
]
]
- | apply rule (⇒_e (A∧¬B⇒C) (A∧¬B));
+ | apply rule (⇒#e (A∧¬B⇒C) (A∧¬B));
[ apply rule (discharge [h1]);
- | apply rule (∧_i (A) (¬B));
- [ apply rule (⇒_e (D⇒A) (D));
+ | apply rule (∧#i (A) (¬B));
+ [ apply rule (⇒#e (D⇒A) (D));
[ apply rule (discharge [h3]);
| apply rule (discharge [h4]);
]
theorem ex3: (F ⇒ G∨E) ⇒ (G ⇒ ¬L∨E) ⇒ (L⇒F) ⇒ L ⇒ E.
apply rule (prove ((F ⇒ G∨E) ⇒ (G ⇒ ¬L∨E) ⇒ (L⇒F) ⇒ L ⇒ E));
(*BEGIN*)
-apply rule (⇒_i [h1] ((G ⇒ ¬L∨E) ⇒ (L⇒F) ⇒ L ⇒ E));
-apply rule (⇒_i [h2] ((L⇒F) ⇒ L ⇒ E));
-apply rule (⇒_i [h3] (L ⇒ E));
-apply rule (⇒_i [h4] (E));
-apply rule (∨_e (G∨E) [h5] (E) [h6] (E));
- [ apply rule (⇒_e (F ⇒ G∨E) (F));
+apply rule (⇒#i [h1] ((G ⇒ ¬L∨E) ⇒ (L⇒F) ⇒ L ⇒ E));
+apply rule (⇒#i [h2] ((L⇒F) ⇒ L ⇒ E));
+apply rule (⇒#i [h3] (L ⇒ E));
+apply rule (⇒#i [h4] (E));
+apply rule (∨#e (G∨E) [h5] (E) [h6] (E));
+ [ apply rule (⇒#e (F ⇒ G∨E) (F));
[ apply rule (discharge [h1]);
- | apply rule (⇒_e (L⇒F) (L));
+ | apply rule (⇒#e (L⇒F) (L));
[ apply rule (discharge [h3]);
| apply rule (discharge [h4]);
]
]
- |apply rule (∨_e (¬L∨E) [h7] (E) [h8] (E));
- [ apply rule (⇒_e (G⇒¬L∨E) (G));
+ |apply rule (∨#e (¬L∨E) [h7] (E) [h8] (E));
+ [ apply rule (⇒#e (G⇒¬L∨E) (G));
[ apply rule (discharge [h2]);
| apply rule (discharge [h5]);
]
- | apply rule (⊥_e (⊥));
- apply rule (¬_e (¬L) (L));
+ | apply rule (⊥#e (⊥));
+ apply rule (¬#e (¬L) (L));
[ apply rule (discharge [h7]);
| apply rule (discharge [h4]);
]
theorem ex4: ¬(A∧B) ⇒ ¬A∨¬B.
apply rule (prove (¬(A∧B) ⇒ ¬A∨¬B));
(*BEGIN*)
-apply rule (⇒_i [h1] (¬A∨¬B));
-apply rule (∨_e (A ∨ ¬A) [h2] ((¬A∨¬B)) [h3] ((¬A∨¬B)));
+apply rule (⇒#i [h1] (¬A∨¬B));
+apply rule (∨#e (A ∨ ¬A) [h2] ((¬A∨¬B)) [h3] ((¬A∨¬B)));
[ apply rule (lem 0 EM);
- | apply rule (∨_e (B ∨ ¬B) [h4] ((¬A∨¬B)) [h5] ((¬A∨¬B)));
+ | apply rule (∨#e (B ∨ ¬B) [h4] ((¬A∨¬B)) [h5] ((¬A∨¬B)));
[ apply rule (lem 0 EM);
- | apply rule (⊥_e (⊥));
- apply rule (¬_e (¬(A∧B)) (A∧B));
+ | apply rule (⊥#e (⊥));
+ apply rule (¬#e (¬(A∧B)) (A∧B));
[ apply rule (discharge [h1]);
- | apply rule (∧_i (A) (B));
+ | apply rule (∧#i (A) (B));
[ apply rule (discharge [h2]);
| apply rule (discharge [h4]);
]
]
- | apply rule (∨_i_r (¬B));
+ | apply rule (∨#i_r (¬B));
apply rule (discharge [h5]);
]
- | apply rule (∨_i_l (¬A));
+ | apply rule (∨#i_l (¬A));
apply rule (discharge [h3]);
]
(*END*)
theorem ex5: ¬(A∨B) ⇒ (¬A ∧ ¬B).
apply rule (prove (¬(A∨B) ⇒ (¬A ∧ ¬B)));
(*BEGIN*)
-apply rule (⇒_i [h1] (¬A ∧ ¬B));
-apply rule (∨_e (A∨¬A) [h2] (¬A ∧ ¬B) [h3] (¬A ∧ ¬B));
+apply rule (⇒#i [h1] (¬A ∧ ¬B));
+apply rule (∨#e (A∨¬A) [h2] (¬A ∧ ¬B) [h3] (¬A ∧ ¬B));
[ apply rule (lem 0 EM);
- | apply rule (⊥_e (⊥));
- apply rule (¬_e (¬(A∨B)) (A∨B));
+ | apply rule (⊥#e (⊥));
+ apply rule (¬#e (¬(A∨B)) (A∨B));
[ apply rule (discharge [h1]);
- | apply rule (∨_i_l (A));
+ | apply rule (∨#i_l (A));
apply rule (discharge [h2]);
]
- | apply rule (∨_e (B∨¬B) [h10] (¬A ∧ ¬B) [h11] (¬A ∧ ¬B));
+ | apply rule (∨#e (B∨¬B) [h10] (¬A ∧ ¬B) [h11] (¬A ∧ ¬B));
[ apply rule (lem 0 EM);
- | apply rule (⊥_e (⊥));
- apply rule (¬_e (¬(A∨B)) (A∨B));
+ | apply rule (⊥#e (⊥));
+ apply rule (¬#e (¬(A∨B)) (A∨B));
[ apply rule (discharge [h1]);
- | apply rule (∨_i_r (B));
+ | apply rule (∨#i_r (B));
apply rule (discharge [h10]);
]
- | apply rule (∧_i (¬A) (¬B));
+ | apply rule (∧#i (¬A) (¬B));
[ apply rule (discharge [h3]);
|apply rule (discharge [h11]);
]
theorem ex6: ¬A ∧ ¬B ⇒ ¬(A∨B).
apply rule (prove (¬A ∧ ¬B ⇒ ¬(A∨B)));
(*BEGIN*)
-apply rule (⇒_i [h1] (¬(A∨B)));
-apply rule (¬_i [h2] (⊥));
-apply rule (∨_e (A∨B) [h3] (⊥) [h3] (⊥));
+apply rule (⇒#i [h1] (¬(A∨B)));
+apply rule (¬#i [h2] (⊥));
+apply rule (∨#e (A∨B) [h3] (⊥) [h3] (⊥));
[ apply rule (discharge [h2]);
- | apply rule (¬_e (¬A) (A));
- [ apply rule (∧_e_l (¬A∧¬B));
+ | apply rule (¬#e (¬A) (A));
+ [ apply rule (∧#e_l (¬A∧¬B));
apply rule (discharge [h1]);
| apply rule (discharge [h3]);
]
- | apply rule (¬_e (¬B) (B));
- [ apply rule (∧_e_r (¬A∧¬B));
+ | apply rule (¬#e (¬B) (B));
+ [ apply rule (∧#e_r (¬A∧¬B));
apply rule (discharge [h1]);
| apply rule (discharge [h3]);
]
theorem ex7: ((A ⇒ ⊥) ⇒ ⊥) ⇒ A.
apply rule (prove (((A ⇒ ⊥) ⇒ ⊥) ⇒ A));
(*BEGIN*)
-apply rule (⇒_i [h1] (A));
+apply rule (⇒#i [h1] (A));
apply rule (RAA [h2] (⊥));
-apply rule (⇒_e ((A⇒⊥)⇒⊥) (A⇒⊥));
+apply rule (⇒#e ((A⇒⊥)⇒⊥) (A⇒⊥));
[ apply rule (discharge [h1]);
- | apply rule (⇒_i [h3] (⊥));
- apply rule (¬_e (¬A) (A));
+ | apply rule (⇒#i [h3] (⊥));
+ apply rule (¬#e (¬A) (A));
[ apply rule (discharge [h2]);
| apply rule (discharge [h3]);
]
assume A: CProp.
apply rule (prove (A ∨ ¬A));
apply rule (RAA [H] (⊥)).
-apply rule (¬_e (¬(A ∨ ¬A)) (A ∨ ¬A));
+apply rule (¬#e (¬(A ∨ ¬A)) (A ∨ ¬A));
[ apply rule (discharge [H]).
- | apply rule (⊥_e (⊥));
- apply rule (¬_e (¬¬A) (¬A));
- [ apply rule (¬_i [K] (⊥)).
- apply rule (¬_e (¬(A ∨ ¬A)) (A ∨ ¬A));
+ | apply rule (⊥#e (⊥));
+ apply rule (¬#e (¬¬A) (¬A));
+ [ apply rule (¬#i [K] (⊥)).
+ apply rule (¬#e (¬(A ∨ ¬A)) (A ∨ ¬A));
[ apply rule (discharge [H]).
- | apply rule (∨_i_r (¬A)).
+ | apply rule (∨#i_r (¬A)).
apply rule (discharge [K]).
]
- | apply rule (¬_i [K] (⊥)).
- apply rule (¬_e (¬(A ∨ ¬A)) (A ∨ ¬A));
+ | apply rule (¬#i [K] (⊥)).
+ apply rule (¬#e (¬(A ∨ ¬A)) (A ∨ ¬A));
[ apply rule (discharge [H]).
- | apply rule (∨_i_l (A)).
+ | apply rule (∨#i_l (A)).
apply rule (discharge [K]).
]
]
theorem ex0 : (F∧¬G ⇒ L ⇒ E) ⇒ (G ∧ L ⇒ E) ⇒ F ∧ L ⇒ E.
apply rule (prove ((F∧¬G ⇒ L ⇒ E) ⇒ (G ∧ L ⇒ E) ⇒ F ∧ L ⇒ E));
(*BEGIN*)
-apply rule (⇒_i [h1] ((G ∧ L ⇒ E) ⇒ F ∧ L ⇒ E));
-apply rule (⇒_i [h2] (F ∧ L ⇒ E));
-apply rule (⇒_i [h3] (E));
-apply rule (∨_e (G∨¬G) [h4] (E) [h4] (E));
+apply rule (⇒#i [h1] ((G ∧ L ⇒ E) ⇒ F ∧ L ⇒ E));
+apply rule (⇒#i [h2] (F ∧ L ⇒ E));
+apply rule (⇒#i [h3] (E));
+apply rule (∨#e (G∨¬G) [h4] (E) [h4] (E));
[ apply rule (lem 0 EM);
- | apply rule (⇒_e (G∧L⇒E) (G∧L));
+ | apply rule (⇒#e (G∧L⇒E) (G∧L));
[ apply rule (discharge [h2]);
- | apply rule (∧_i (G) (L));
+ | apply rule (∧#i (G) (L));
[ apply rule (discharge [h4]);
- | apply rule (∧_e_r (F∧L));
+ | apply rule (∧#e_r (F∧L));
apply rule (discharge [h3]);
]
]
- | apply rule (⇒_e (L⇒E) (L));
- [ apply rule (⇒_e (F∧¬G ⇒ L ⇒ E) (F∧¬G));
+ | apply rule (⇒#e (L⇒E) (L));
+ [ apply rule (⇒#e (F∧¬G ⇒ L ⇒ E) (F∧¬G));
[ apply rule (discharge [h1]);
- | apply rule (∧_i (F) (¬G));
- [ apply rule (∧_e_l (F∧L));
+ | apply rule (∧#i (F) (¬G));
+ [ apply rule (∧#e_l (F∧L));
apply rule (discharge [h3]);
| apply rule (discharge [h4]);
]
]
- | apply rule (∧_e_r (F∧L));
+ | apply rule (∧#e_r (F∧L));
apply rule (discharge [h3]);
]
]
theorem ex1 : (F ⇒ G∨E) ⇒ (G ⇒ ¬ (L ∧ ¬E)) ⇒ (¬L ∨ F) ⇒ L ⇒ E.
apply rule (prove (F ⇒ G∨E) ⇒ (G ⇒ ¬ (L ∧ ¬E)) ⇒ (¬L ∨ F) ⇒ L ⇒ E);
(*BEGIN*)
-apply rule (⇒_i [h1] ((G ⇒ ¬ (L ∧ ¬E)) ⇒ (¬L ∨ F) ⇒ L ⇒ E));
-apply rule (⇒_i [h2] ((¬L ∨ F) ⇒ L ⇒ E));
-apply rule (⇒_i [h3] (L ⇒ E));
-apply rule (⇒_i [h4] (E));
-apply rule (∨_e (¬L∨F) [h5] (E) [h5] (E));
+apply rule (⇒#i [h1] ((G ⇒ ¬ (L ∧ ¬E)) ⇒ (¬L ∨ F) ⇒ L ⇒ E));
+apply rule (⇒#i [h2] ((¬L ∨ F) ⇒ L ⇒ E));
+apply rule (⇒#i [h3] (L ⇒ E));
+apply rule (⇒#i [h4] (E));
+apply rule (∨#e (¬L∨F) [h5] (E) [h5] (E));
[ apply rule (discharge [h3]);
- | apply rule (⊥_e (⊥));
- apply rule (¬_e (¬L) (L));
+ | apply rule (⊥#e (⊥));
+ apply rule (¬#e (¬L) (L));
[ apply rule (discharge [h5]);
| apply rule (discharge [h4]);
]
- | apply rule (∨_e (G∨E) [h6] (E) [h6] (E));
- [ apply rule (⇒_e (F⇒G∨E) (F));
+ | apply rule (∨#e (G∨E) [h6] (E) [h6] (E));
+ [ apply rule (⇒#e (F⇒G∨E) (F));
[ apply rule (discharge [h1]);
| apply rule (discharge [h5]);
]
| apply rule (RAA [h7] (⊥));
- apply rule (¬_e (¬ (L ∧ ¬E)) (L ∧ ¬E));
- [ apply rule (⇒_e (G ⇒ ¬ (L ∧ ¬E)) (G));
+ apply rule (¬#e (¬ (L ∧ ¬E)) (L ∧ ¬E));
+ [ apply rule (⇒#e (G ⇒ ¬ (L ∧ ¬E)) (G));
[apply rule (discharge [h2]);
|apply rule (discharge [h6]);
]
- | apply rule (∧_i (L) (¬E));
+ | apply rule (∧#i (L) (¬E));
[ apply rule (discharge [h4]);
| apply rule (discharge [h7]);
]
theorem ex2 : (F∧G ⇒ E) ⇒ (H ⇒ E∨F) ⇒ (¬G ⇒ ¬H) ⇒ H ⇒ E.
apply rule (prove ((F∧G ⇒ E) ⇒ (H ⇒ E∨F) ⇒ (¬G ⇒ ¬H) ⇒ H ⇒ E));
(*BEGIN*)
-apply rule (⇒_i [h1] ((H ⇒ E∨F) ⇒ (¬G ⇒ ¬H) ⇒ H ⇒ E));
-apply rule (⇒_i [h2] ((¬G ⇒ ¬H) ⇒ H ⇒ E));
-apply rule (⇒_i [h3] (H ⇒ E));
-apply rule (⇒_i [h4] (E));
-apply rule (∨_e (G∨¬G) [h5] (E) [h5] (E));
+apply rule (⇒#i [h1] ((H ⇒ E∨F) ⇒ (¬G ⇒ ¬H) ⇒ H ⇒ E));
+apply rule (⇒#i [h2] ((¬G ⇒ ¬H) ⇒ H ⇒ E));
+apply rule (⇒#i [h3] (H ⇒ E));
+apply rule (⇒#i [h4] (E));
+apply rule (∨#e (G∨¬G) [h5] (E) [h5] (E));
[apply rule (lem 0 EM);
- | apply rule (∨_e (E∨F) [h6] (E) [h6] (E));
- [ apply rule (⇒_e (H ⇒ E∨F) (H));
+ | apply rule (∨#e (E∨F) [h6] (E) [h6] (E));
+ [ apply rule (⇒#e (H ⇒ E∨F) (H));
[ apply rule (discharge [h2]);
| apply rule (discharge [h4]);
]
| apply rule (discharge [h6]);
- | apply rule (⇒_e (F∧G ⇒ E) (F∧G));
+ | apply rule (⇒#e (F∧G ⇒ E) (F∧G));
[apply rule (discharge [h1]);
- |apply rule (∧_i (F) (G));
+ |apply rule (∧#i (F) (G));
[ apply rule (discharge [h6]);
| apply rule (discharge [h5]);
]
]
]
- | apply rule (⊥_e (⊥));
- apply rule (¬_e (¬H) (H));
- [ apply rule (⇒_e (¬G ⇒ ¬H) (¬G));
+ | apply rule (⊥#e (⊥));
+ apply rule (¬#e (¬H) (H));
+ [ apply rule (⇒#e (¬G ⇒ ¬H) (¬G));
[ apply rule (discharge [h3]);
| apply rule (discharge [h5]);
]
theorem ex3 : (F∧G ⇒ E) ⇒ (¬E ⇒ G∨¬L) ⇒ (L ⇒ F) ⇒ L ⇒ E.
apply rule (prove ((F∧G ⇒ E) ⇒ (¬E ⇒ G∨¬L) ⇒ (L ⇒ F) ⇒ L ⇒ E)).
(*BEGIN*)
-apply rule (⇒_i [h1] ((¬E ⇒ G∨¬L) ⇒ (L ⇒ F) ⇒ L ⇒ E));
-apply rule (⇒_i [h2] ((L ⇒ F) ⇒ L ⇒ E));
-apply rule (⇒_i [h3] (L ⇒ E));
-apply rule (⇒_i [h4] (E));
+apply rule (⇒#i [h1] ((¬E ⇒ G∨¬L) ⇒ (L ⇒ F) ⇒ L ⇒ E));
+apply rule (⇒#i [h2] ((L ⇒ F) ⇒ L ⇒ E));
+apply rule (⇒#i [h3] (L ⇒ E));
+apply rule (⇒#i [h4] (E));
apply rule (RAA [h5] (⊥));
-apply rule (∨_e (G∨¬L) [h6] (⊥) [h6] (⊥));
- [ apply rule (⇒_e (¬E ⇒ G∨¬L) (¬E));
+apply rule (∨#e (G∨¬L) [h6] (⊥) [h6] (⊥));
+ [ apply rule (⇒#e (¬E ⇒ G∨¬L) (¬E));
[ apply rule (discharge [h2]);
| apply rule (discharge [h5]);
]
- | apply rule (¬_e (¬E) (E));
+ | apply rule (¬#e (¬E) (E));
[ apply rule (discharge [h5]);
- | apply rule (⇒_e (F∧G ⇒ E) (F∧G));
+ | apply rule (⇒#e (F∧G ⇒ E) (F∧G));
[ apply rule (discharge [h1]);
- | apply rule (∧_i (F) (G));
- [ apply rule (⇒_e (L⇒F) (L));
+ | apply rule (∧#i (F) (G));
+ [ apply rule (⇒#e (L⇒F) (L));
[ apply rule (discharge [h3]);
| apply rule (discharge [h4]);
]
]
]
]
- | apply rule (¬_e (¬L) (L));
+ | apply rule (¬#e (¬L) (L));
[ apply rule (discharge [h6]);
| apply rule (discharge [h4]);
]
con premesse multiple sono seguite da `[`, `|` e `]`.
Ad esempio:
- apply rule (∧_i (A∨B) ⊥);
+ apply rule (∧#i (A∨B) ⊥);
[ …continua qui il sottoalbero per (A∨B)
| …continua qui il sottoalbero per ⊥
]
scrivete sopra la linea che rappresenta l'applicazione della
regola stessa. Ad esempio:
- aply rule (∨_e (A∨B) [h1] C [h2] C);
+ aply rule (∨#e (A∨B) [h1] C [h2] C);
[ …prova di (A∨B)
| …prova di C sotto l'ipotesi A (di nome h1)
| …prova di C sotto l'ipotesi B (di nome h2)
assume A: CProp.
apply rule (prove (A ∨ ¬A));
apply rule (RAA [H] (⊥)).
-apply rule (¬_e (¬(A ∨ ¬A)) (A ∨ ¬A));
+apply rule (¬#e (¬(A ∨ ¬A)) (A ∨ ¬A));
[ apply rule (discharge [H]).
- | apply rule (⊥_e (⊥));
- apply rule (¬_e (¬¬A) (¬A));
- [ apply rule (¬_i [K] (⊥)).
- apply rule (¬_e (¬(A ∨ ¬A)) (A ∨ ¬A));
+ | apply rule (⊥#e (⊥));
+ apply rule (¬#e (¬¬A) (¬A));
+ [ apply rule (¬#i [K] (⊥)).
+ apply rule (¬#e (¬(A ∨ ¬A)) (A ∨ ¬A));
[ apply rule (discharge [H]).
- | apply rule (∨_i_r (¬A)).
+ | apply rule (∨#i_r (¬A)).
apply rule (discharge [K]).
]
- | apply rule (¬_i [K] (⊥)).
- apply rule (¬_e (¬(A ∨ ¬A)) (A ∨ ¬A));
+ | apply rule (¬#i [K] (⊥)).
+ apply rule (¬#e (¬(A ∨ ¬A)) (A ∨ ¬A));
[ apply rule (discharge [H]).
- | apply rule (∨_i_l (A)).
+ | apply rule (∨#i_l (A)).
apply rule (discharge [K]).
]
]
lemma ex1: ¬(∃x.P x) ⇒ ∀x.¬ P x.
apply rule (prove (¬(∃x.P x) ⇒ ∀x.¬ P x));
(*BEGIN*)
-apply rule (⇒_i [h1] (∀x.¬ P x));
-apply rule (∀_i {l} (¬P l));
-apply rule (¬_i [h2] (⊥));
-apply rule (¬_e (¬(∃x.P x)) (∃x.P x));
+apply rule (⇒#i [h1] (∀x.¬ P x));
+apply rule (∀#i {l} (¬P l));
+apply rule (¬#i [h2] (⊥));
+apply rule (¬#e (¬(∃x.P x)) (∃x.P x));
[ apply rule (discharge [h1]);
- | apply rule (∃_i {l} (P l));
+ | apply rule (∃#i {l} (P l));
apply rule (discharge [h2]);
]
(*END*)
lemma ex2: ¬(∀x.P x) ⇒ ∃x.¬ P x.
apply rule (prove (¬(∀x.P x) ⇒ ∃x.¬ P x));
(*BEGIN*)
-apply rule (⇒_i [h1] (∃x.¬ P x));
+apply rule (⇒#i [h1] (∃x.¬ P x));
apply rule (RAA [h2] (⊥));
-apply rule (¬_e (¬(∀x.P x)) (∀x.P x));
+apply rule (¬#e (¬(∀x.P x)) (∀x.P x));
[ apply rule (discharge [h2]);
- | apply rule (∀_i {y} (P y));
+ | apply rule (∀#i {y} (P y));
apply rule (RAA [h3] (⊥));
- apply rule (¬_e (¬∃x.¬ P x) (∃x.¬ P x));
+ apply rule (¬#e (¬∃x.¬ P x) (∃x.¬ P x));
[ apply rule (discharge [h2]);
- | apply rule (∃_i {y} (¬P y));
+ | apply rule (∃#i {y} (¬P y));
apply rule (discharge [h3]);
]
]
lemma ex3: ((∃x.P x) ⇒ Q c) ⇒ ∀x.P x ⇒ Q c.
apply rule (prove (((∃x.P x) ⇒ Q c) ⇒ ∀x.P x ⇒ Q c));
(*BEGIN*)
-apply rule (⇒_i [h1] (∀x.P x ⇒ Q c));
-apply rule (∀_i {l} (P l ⇒ Q c));
-apply rule (⇒_i [h2] (Q c));
-apply rule (⇒_e ((∃x.P x) ⇒ Q c) (∃x.P x));
+apply rule (⇒#i [h1] (∀x.P x ⇒ Q c));
+apply rule (∀#i {l} (P l ⇒ Q c));
+apply rule (⇒#i [h2] (Q c));
+apply rule (⇒#e ((∃x.P x) ⇒ Q c) (∃x.P x));
[ apply rule (discharge [h1]);
- | apply rule (∃_i {l} (P l));
+ | apply rule (∃#i {l} (P l));
apply rule (discharge [h2]);
]
(*END*)
lemma ex4: ((∀x.P x) ⇒ Q c) ⇒ ∃x.P x ⇒ Q c.
apply rule (prove (((∀x.P x) ⇒ Q c) ⇒ ∃x.P x ⇒ Q c));
(*BEGIN*)
-apply rule (⇒_i [h1] (∃x.P x ⇒ Q c));
-apply rule (∨_e ((∀x.P x) ∨ ¬(∀x.P x)) [h3] (?) [h3] (?));
+apply rule (⇒#i [h1] (∃x.P x ⇒ Q c));
+apply rule (∨#e ((∀x.P x) ∨ ¬(∀x.P x)) [h3] (?) [h3] (?));
[ apply rule (lem 0 EM);
- | apply rule (∃_i {y} (P y ⇒ Q c));
- apply rule (⇒_i [h4] (Q c));
- apply rule (⇒_e ((∀x.P x) ⇒ Q c) ((∀x.P x)));
+ | apply rule (∃#i {y} (P y ⇒ Q c));
+ apply rule (⇒#i [h4] (Q c));
+ apply rule (⇒#e ((∀x.P x) ⇒ Q c) ((∀x.P x)));
[ apply rule (discharge [h1]);
| apply rule (discharge [h3]);
]
- | apply rule (∃_e (∃x.¬P x) {y} [h4] (∃x.P x ⇒ Q c));
+ | apply rule (∃#e (∃x.¬P x) {y} [h4] (∃x.P x ⇒ Q c));
[ apply rule (lem 1 ex2 (¬(∀x.P x)));
apply rule (discharge [h3]);
- | apply rule (∃_i {y} (P y ⇒ Q c));
- apply rule (⇒_i [h5] (Q c));
- apply rule (⊥_e (⊥));
- apply rule (¬_e (¬P y) (P y));
+ | apply rule (∃#i {y} (P y ⇒ Q c));
+ apply rule (⇒#i [h5] (Q c));
+ apply rule (⊥#e (⊥));
+ apply rule (¬#e (¬P y) (P y));
[ apply rule (discharge [h4]);
| apply rule (discharge [h5]);
]
con premesse multiple sono seguite da `[`, `|` e `]`.
Ad esempio:
- apply rule (∧_i (A∨B) ⊥);
+ apply rule (∧#i (A∨B) ⊥);
[ …continua qui il sottoalbero per (A∨B)
| …continua qui il sottoalbero per ⊥
]
scrivete sopra la linea che rappresenta l'applicazione della
regola stessa. Ad esempio:
- apply rule (∨_e (A∨B) [h1] C [h2] C);
+ apply rule (∨#e (A∨B) [h1] C [h2] C);
[ …prova di (A∨B)
| …prova di C sotto l'ipotesi A (di nome h1)
| …prova di C sotto l'ipotesi B (di nome h2)
l'unica forma di `apply rule lem` che potete usare è
`apply rule (lem 0 nome_assioma)` *)
(*BEGIN*)
-apply rule (⇒_e ((∀n.(n + O) = n ⇒ ((S n) + O) = (S n)) ⇒ (∀x.(x + O) = x)) (∀n.(n + O) = n ⇒ ((S n) + O) = (S n)));
- [ apply rule (⇒_e ((O + O) = O ⇒ (∀n.n + O = n ⇒ ((S n) + O) = (S n)) ⇒ (∀x.(x + O) = x)) (O + O = O));
+apply rule (⇒#e ((∀n.(n + O) = n ⇒ ((S n) + O) = (S n)) ⇒ (∀x.(x + O) = x)) (∀n.(n + O) = n ⇒ ((S n) + O) = (S n)));
+ [ apply rule (⇒#e ((O + O) = O ⇒ (∀n.n + O = n ⇒ ((S n) + O) = (S n)) ⇒ (∀x.(x + O) = x)) (O + O = O));
[ apply rule (lem 0 induct);
- | apply rule (∀_e {O} (∀y.(O + y) =y));
+ | apply rule (∀#e {O} (∀y.(O + y) =y));
apply rule (lem 0 plus_O);
]
- | apply rule (∀_i {n} (n + O = n ⇒ ((S n) + O) = (S n)));
- apply rule (⇒_i [H] (((S n) + O) = (S n)));
- apply rule (⇒_e ((S (n + O)) = (S n) ⇒ ((S n) + O) = (S n)) ((S (n + O)) = (S n)));
- [ apply rule (⇒_e (((S n) + O) =(S (n + O)) ⇒ (S (n + O)) = (S n)⇒((S n) + O) =(S n)) (((S n) + O) =(S (n + O))));
- [ apply rule (∀_e {S n} (∀z.((S n) + O) =(S (n + O))⇒(S (n + O))= z⇒((S n) + O) =z));
- apply rule (∀_e {(S (n + O))} (∀y.∀z.((S n) + O) =y ⇒ y = z ⇒ ((S n) + O) =z));
- apply rule (∀_e {(S n) + O} (∀x.∀y.∀z.x = y ⇒ y = z ⇒ x = z));
+ | apply rule (∀#i {n} (n + O = n ⇒ ((S n) + O) = (S n)));
+ apply rule (⇒#i [H] (((S n) + O) = (S n)));
+ apply rule (⇒#e ((S (n + O)) = (S n) ⇒ ((S n) + O) = (S n)) ((S (n + O)) = (S n)));
+ [ apply rule (⇒#e (((S n) + O) =(S (n + O)) ⇒ (S (n + O)) = (S n)⇒((S n) + O) =(S n)) (((S n) + O) =(S (n + O))));
+ [ apply rule (∀#e {S n} (∀z.((S n) + O) =(S (n + O))⇒(S (n + O))= z⇒((S n) + O) =z));
+ apply rule (∀#e {(S (n + O))} (∀y.∀z.((S n) + O) =y ⇒ y = z ⇒ ((S n) + O) =z));
+ apply rule (∀#e {(S n) + O} (∀x.∀y.∀z.x = y ⇒ y = z ⇒ x = z));
apply rule (lem 0 trans);
- | apply rule (∀_e {O} (∀m.(S n) + m = (S (n + m))));
- apply rule (∀_e {n} (∀n.∀m.(S n) + m = (S (n + m))));
+ | apply rule (∀#e {O} (∀m.(S n) + m = (S (n + m))));
+ apply rule (∀#e {n} (∀n.∀m.(S n) + m = (S (n + m))));
apply rule (lem 0 plus_S);
]
- | apply rule (⇒_e (n + O = n ⇒ (S (n + O)) = (S n)) (n + O = n));
- [ apply rule (∀_e {n} (∀w.n + O = w ⇒ (S (n + O)) = (S w)));
- apply rule (∀_e {(n + O)} (∀y.∀w.y = w ⇒ (S y) = (S w)));
+ | apply rule (⇒#e (n + O = n ⇒ (S (n + O)) = (S n)) (n + O = n));
+ [ apply rule (∀#e {n} (∀w.n + O = w ⇒ (S (n + O)) = (S w)));
+ apply rule (∀#e {(n + O)} (∀y.∀w.y = w ⇒ (S y) = (S w)));
apply rule (lem 0 eq_S);
| apply rule (discharge [H]);
]
definition if_then_else ≝ λT:Type.λe,t,f.match e return λ_.T with [ true ⇒ t | false ⇒ f].
notation > "'if' term 19 e 'then' term 19 t 'else' term 90 f" non associative with precedence 19 for @{ 'if_then_else $e $t $f }.
notation < "'if' \nbsp term 19 e \nbsp 'then' \nbsp term 19 t \nbsp 'else' \nbsp term 90 f \nbsp" non associative with precedence 19 for @{ 'if_then_else $e $t $f }.
-interpretation "Formula if_then_else" 'if_then_else e t f = (if_then_else _ e t f).
+interpretation "Formula if_then_else" 'if_then_else e t f = (if_then_else ? e t f).
definition max ≝ λn,m. if eqb (n - m) 0 then m else n.
definition min ≝ λn,m. if eqb (n - m) 0 then n else m.
(* Ripasso 2
=========
- La semantica di una formula `F` in un mondo `v`: `[[ F ]]_v`
+ La semantica di una formula `F` in un mondo `v`: `[[ F ]]v`
*)
let rec sem (v: nat → nat) (F: Formula) on F : nat ≝
match F with
*)
notation < "[[ \nbsp term 19 a \nbsp ]] \nbsp \sub term 90 v" non associative with precedence 90 for @{ 'semantics $v $a }.
notation > "[[ term 19 a ]] \sub term 90 v" non associative with precedence 90 for @{ 'semantics $v $a }.
-notation > "[[ term 19 a ]]_ term 90 v" non associative with precedence 90 for @{ sem $v $a }.
+notation > "[[ term 19 a ]] term 90 v" non associative with precedence 90 for @{ sem $v $a }.
interpretation "Semantic of Formula" 'semantics v a = (sem v a).
-lemma sem_bool : ∀F,v. [[ F ]]_v = 0 ∨ [[ F ]]_v = 1. intros; elim F; simplify; [left;reflexivity; |right;reflexivity; |cases (v n);[left;|cases n1;right;]reflexivity; |4,5,6: cases H; cases H1; rewrite > H2; rewrite > H3; simplify; first [ left;reflexivity | right; reflexivity ]. |cases H; rewrite > H1; simplify;[right|left]reflexivity;] qed.
+lemma sem_bool : ∀F,v. [[ F ]]v = 0 ∨ [[ F ]]v = 1. intros; elim F; simplify; [left;reflexivity; |right;reflexivity; |cases (v n);[left;|cases n1;right;]reflexivity; |4,5,6: cases H; cases H1; rewrite > H2; rewrite > H3; simplify; first [ left;reflexivity | right; reflexivity ]. |cases H; rewrite > H1; simplify;[right|left]reflexivity;] qed.
(* Ripasso 3
=========
notation < "t [ \nbsp term 19 a / term 19 b \nbsp ]" non associative with precedence 19 for @{ 'substitution $b $a $t }.
notation > "t [ term 90 a / term 90 b]" non associative with precedence 19 for @{ 'substitution $b $a $t }.
interpretation "Substitution for Formula" 'substitution b a t = (subst b a t).
-definition equiv ≝ λF1,F2. ∀v.[[ F1 ]]_v = [[ F2 ]]_v.
+definition equiv ≝ λF1,F2. ∀v.[[ F1 ]]v = [[ F2 ]]v.
notation "hvbox(a \nbsp break mstyle color #0000ff (≡) \nbsp b)" non associative with precedence 45 for @{ 'equivF $a $b }.
notation > "a ≡ b" non associative with precedence 50 for @{ equiv $a $b }.
interpretation "equivalence for Formulas" 'equivF a b = (equiv a b).
-lemma min_1_sem: ∀F,v.min 1 [[ F ]]_v = [[ F ]]_v. intros; cases (sem_bool F v); rewrite > H; reflexivity; qed.
-lemma max_0_sem: ∀F,v.max [[ F ]]_v 0 = [[ F ]]_v. intros; cases (sem_bool F v); rewrite > H; reflexivity; qed.
+lemma min_1_sem: ∀F,v.min 1 [[ F ]]v = [[ F ]]v. intros; cases (sem_bool F v); rewrite > H; reflexivity; qed.
+lemma max_0_sem: ∀F,v.max [[ F ]]v 0 = [[ F ]]v. intros; cases (sem_bool F v); rewrite > H; reflexivity; qed.
definition IFTE := λA,B,C:Formula. FOr (FAnd A B) (FAnd (FNot A) C).
(*DOCBEGIN
Per portare a termine l'esercitazione sono necessari i seguenti lemmi:
* lemma `decidable_eq_nat` : `∀x,y.x = y ∨ x ≠ y`
-* lemma `sem_bool` : `∀F,v. [[ F ]]_v = 0 ∨ [[ F ]]_v = 1`
+* lemma `sem_bool` : `∀F,v. [[ F ]]v = 0 ∨ [[ F ]]v = 1`
* lemma `not_eq_to_eqb_false` : `∀x,y.x ≠ y → eqb x y = false`
* lemma `eq_to_eqb_true` : `∀x,y.x = y → eqb x y = true`
-* lemma `min_1_sem` : `∀F,v.min 1 [[ F ]]_v = [[ F ]]_v`
-* lemma `max_0_sem` : `∀F,v.max [[ F ]]_v 0 = [[ F ]]_v`
+* lemma `min_1_sem` : `∀F,v.min 1 [[ F ]]v = [[ F ]]v`
+* lemma `max_0_sem` : `∀F,v.max [[ F ]]v 0 = [[ F ]]v`
Nota su `x = y` e `eqb x y`
---------------------------
Vediamo solo la dimostrazione del primo, essendo il secondo del tutto analogo.
Il lemma asserisce quanto segue:
- ∀F,x,v. [[ FAtom x ]]_v = 0 → [[ F[FBot/x] ]]_v = [[ F ]]_v
+ ∀F,x,v. [[ FAtom x ]]v = 0 → [[ F[FBot/x] ]]v = [[ F ]]v
Una volta assunte la formula `F`, l'atomo `x`, il mondo `v` e aver
-supposto che `[[ FAtom x ]]_v = 0` si procede per induzione su `F`.
+supposto che `[[ FAtom x ]]v = 0` si procede per induzione su `F`.
I casi `FTop` e `FBot` sono banali. Nei casi `FAnd/FOr/FImpl/FNot`,
una volta assunte le sottoformule e le relative ipotesi induttive,
si conclude con una catena di uguaglianze.
Entrambi i casi si concludono con una catena di uguaglianze.
Il teorema principale si dimostra utilizzando il lemma `sem_bool` per
-ottenre l'ipotesi `[[ FAtom x ]]_v = 0 ∨ [[ FAtom x ]]_v = 1` su cui
+ottenre l'ipotesi `[[ FAtom x ]]v = 0 ∨ [[ FAtom x ]]v = 1` su cui
si procede poi per casi. Entrambi i casi si concludono con
una catena di uguaglianze che utilizza i lemmi dimostrati in precedenza
e i lemmi `min_1_sem` oppure `max_0_sem`.
DOCEND*)
lemma shannon_false:
- ∀F,x,v. [[ FAtom x ]]_v = 0 → [[ F[FBot/x] ]]_v = [[ F ]]_v.
+ ∀F,x,v. [[ FAtom x ]]v = 0 → [[ F[FBot/x] ]]v = [[ F ]]v.
(*BEGIN*)
assume F : Formula.
assume x : ℕ.
assume v : (ℕ → ℕ).
-suppose ([[ FAtom x ]]_v = 0) (H).
-we proceed by induction on F to prove ([[ F[FBot/x] ]]_v = [[ F ]]_v).
+suppose ([[ FAtom x ]]v = 0) (H).
+we proceed by induction on F to prove ([[ F[FBot/x] ]]v = [[ F ]]v).
case FBot.
- the thesis becomes ([[ FBot[FBot/x] ]]_v = [[ FBot ]]_v).
- the thesis becomes ([[ FBot ]]_v = [[ FBot ]]_v).
+ the thesis becomes ([[ FBot[FBot/x] ]]v = [[ FBot ]]v).
+ the thesis becomes ([[ FBot ]]v = [[ FBot ]]v).
done.
case FTop.
- the thesis becomes ([[ FTop[FBot/x] ]]_v = [[ FTop ]]_v).
- the thesis becomes ([[ FTop ]]_v = [[ FTop ]]_v).
+ the thesis becomes ([[ FTop[FBot/x] ]]v = [[ FTop ]]v).
+ the thesis becomes ([[ FTop ]]v = [[ FTop ]]v).
done.
case FAtom.
assume n : ℕ.
- the thesis becomes ([[ (FAtom n)[FBot/x] ]]_v = [[ FAtom n ]]_v).
- the thesis becomes ([[ if eqb n x then FBot else (FAtom n) ]]_v = [[ FAtom n ]]_v).
+ the thesis becomes ([[ (FAtom n)[FBot/x] ]]v = [[ FAtom n ]]v).
+ the thesis becomes ([[ if eqb n x then FBot else (FAtom n) ]]v = [[ FAtom n ]]v).
by decidable_eq_nat we proved (n = x ∨ n ≠ x) (H1).
- we proceed by cases on H1 to prove ([[ if eqb n x then FBot else (FAtom n) ]]_v = [[ FAtom n ]]_v).
+ we proceed by cases on H1 to prove ([[ if eqb n x then FBot else (FAtom n) ]]v = [[ FAtom n ]]v).
case Left.
by H2, eq_to_eqb_true we proved (eqb n x = true) (H3).
conclude
- ([[ if eqb n x then FBot else (FAtom n) ]]_v)
- = ([[ if true then FBot else (FAtom n) ]]_v) by H3.
- = ([[ FBot ]]_v).
+ ([[ if eqb n x then FBot else (FAtom n) ]]v)
+ = ([[ if true then FBot else (FAtom n) ]]v) by H3.
+ = ([[ FBot ]]v).
= 0.
- = ([[ FAtom x ]]_v) by H.
- = ([[ FAtom n ]]_v) by H2.
+ = ([[ FAtom x ]]v) by H.
+ = ([[ FAtom n ]]v) by H2.
done.
case Right.
by H2, not_eq_to_eqb_false we proved (eqb n x = false) (H3).
conclude
- ([[ if eqb n x then FBot else (FAtom n) ]]_v)
- = ([[ if false then FBot else (FAtom n) ]]_v) by H3.
- = ([[ FAtom n ]]_v).
+ ([[ if eqb n x then FBot else (FAtom n) ]]v)
+ = ([[ if false then FBot else (FAtom n) ]]v) by H3.
+ = ([[ FAtom n ]]v).
done.
case FAnd.
assume f1 : Formula.
- by induction hypothesis we know ([[ f1[FBot/x] ]]_v = [[ f1 ]]_v) (H1).
+ by induction hypothesis we know ([[ f1[FBot/x] ]]v = [[ f1 ]]v) (H1).
assume f2 : Formula.
- by induction hypothesis we know ([[ f2[FBot/x] ]]_v = [[ f2 ]]_v) (H2).
- the thesis becomes ([[ (FAnd f1 f2)[FBot/x] ]]_v = [[ FAnd f1 f2 ]]_v).
+ by induction hypothesis we know ([[ f2[FBot/x] ]]v = [[ f2 ]]v) (H2).
+ the thesis becomes ([[ (FAnd f1 f2)[FBot/x] ]]v = [[ FAnd f1 f2 ]]v).
conclude
- ([[ (FAnd f1 f2)[FBot/x] ]]_v)
- = ([[ FAnd (f1[FBot/x]) (f2[FBot/x]) ]]_v).
- = (min [[ f1[FBot/x] ]]_v [[ f2[FBot/x] ]]_v).
- = (min [[ f1 ]]_v [[ f2[FBot/x] ]]_v) by H1.
- = (min [[ f1 ]]_v [[ f2 ]]_v) by H2.
- = ([[ FAnd f1 f2 ]]_v).
+ ([[ (FAnd f1 f2)[FBot/x] ]]v)
+ = ([[ FAnd (f1[FBot/x]) (f2[FBot/x]) ]]v).
+ = (min [[ f1[FBot/x] ]]v [[ f2[FBot/x] ]]v).
+ = (min [[ f1 ]]v [[ f2[FBot/x] ]]v) by H1.
+ = (min [[ f1 ]]v [[ f2 ]]v) by H2.
+ = ([[ FAnd f1 f2 ]]v).
done.
case FOr.
assume f1 : Formula.
- by induction hypothesis we know ([[ f1[FBot/x] ]]_v = [[ f1 ]]_v) (H1).
+ by induction hypothesis we know ([[ f1[FBot/x] ]]v = [[ f1 ]]v) (H1).
assume f2 : Formula.
- by induction hypothesis we know ([[ f2[FBot/x] ]]_v = [[ f2 ]]_v) (H2).
- the thesis becomes ([[ (FOr f1 f2)[FBot/x] ]]_v = [[ FOr f1 f2 ]]_v).
+ by induction hypothesis we know ([[ f2[FBot/x] ]]v = [[ f2 ]]v) (H2).
+ the thesis becomes ([[ (FOr f1 f2)[FBot/x] ]]v = [[ FOr f1 f2 ]]v).
conclude
- ([[ (FOr f1 f2)[FBot/x] ]]_v)
- = ([[ FOr (f1[FBot/x]) (f2[FBot/x]) ]]_v).
- = (max [[ f1[FBot/x] ]]_v [[ f2[FBot/x] ]]_v).
- = (max [[ f1 ]]_v [[ f2[FBot/x] ]]_v) by H1.
- = (max [[ f1 ]]_v [[ f2 ]]_v) by H2.
- = ([[ FOr f1 f2 ]]_v).
+ ([[ (FOr f1 f2)[FBot/x] ]]v)
+ = ([[ FOr (f1[FBot/x]) (f2[FBot/x]) ]]v).
+ = (max [[ f1[FBot/x] ]]v [[ f2[FBot/x] ]]v).
+ = (max [[ f1 ]]v [[ f2[FBot/x] ]]v) by H1.
+ = (max [[ f1 ]]v [[ f2 ]]v) by H2.
+ = ([[ FOr f1 f2 ]]v).
done.
case FImpl.
assume f1 : Formula.
- by induction hypothesis we know ([[ f1[FBot/x] ]]_v = [[ f1 ]]_v) (H1).
+ by induction hypothesis we know ([[ f1[FBot/x] ]]v = [[ f1 ]]v) (H1).
assume f2 : Formula.
- by induction hypothesis we know ([[ f2[FBot/x] ]]_v = [[ f2 ]]_v) (H2).
- the thesis becomes ([[ (FImpl f1 f2)[FBot/x] ]]_v = [[ FImpl f1 f2 ]]_v).
+ by induction hypothesis we know ([[ f2[FBot/x] ]]v = [[ f2 ]]v) (H2).
+ the thesis becomes ([[ (FImpl f1 f2)[FBot/x] ]]v = [[ FImpl f1 f2 ]]v).
conclude
- ([[ (FImpl f1 f2)[FBot/x] ]]_v)
- = ([[ FImpl (f1[FBot/x]) (f2[FBot/x]) ]]_v).
- = (max (1 - [[ f1[FBot/x] ]]_v) [[ f2[FBot/x] ]]_v).
- = (max (1 - [[ f1 ]]_v) [[ f2[FBot/x] ]]_v) by H1.
- = (max (1 - [[ f1 ]]_v) [[ f2 ]]_v) by H2.
- = ([[ FImpl f1 f2 ]]_v).
+ ([[ (FImpl f1 f2)[FBot/x] ]]v)
+ = ([[ FImpl (f1[FBot/x]) (f2[FBot/x]) ]]v).
+ = (max (1 - [[ f1[FBot/x] ]]v) [[ f2[FBot/x] ]]v).
+ = (max (1 - [[ f1 ]]v) [[ f2[FBot/x] ]]v) by H1.
+ = (max (1 - [[ f1 ]]v) [[ f2 ]]v) by H2.
+ = ([[ FImpl f1 f2 ]]v).
done.
case FNot.
assume f : Formula.
- by induction hypothesis we know ([[ f[FBot/x] ]]_v = [[ f ]]_v) (H1).
- the thesis becomes ([[ (FNot f)[FBot/x] ]]_v = [[ FNot f ]]_v).
+ by induction hypothesis we know ([[ f[FBot/x] ]]v = [[ f ]]v) (H1).
+ the thesis becomes ([[ (FNot f)[FBot/x] ]]v = [[ FNot f ]]v).
conclude
- ([[ (FNot f)[FBot/x] ]]_v)
- = ([[ FNot (f[FBot/x]) ]]_v).
- = (1 - [[ f[FBot/x] ]]_v).
- = (1 - [[ f ]]_v) by H1.
- = ([[ FNot f ]]_v).
+ ([[ (FNot f)[FBot/x] ]]v)
+ = ([[ FNot (f[FBot/x]) ]]v).
+ = (1 - [[ f[FBot/x] ]]v).
+ = (1 - [[ f ]]v) by H1.
+ = ([[ FNot f ]]v).
done.
(*END*)
qed.
lemma shannon_true:
- ∀F,x,v. [[ FAtom x ]]_v = 1 → [[ F[FTop/x] ]]_v = [[ F ]]_v.
+ ∀F,x,v. [[ FAtom x ]]v = 1 → [[ F[FTop/x] ]]v = [[ F ]]v.
(*BEGIN*)
assume F : Formula.
assume x : ℕ.
assume v : (ℕ → ℕ).
-suppose ([[ FAtom x ]]_v = 1) (H).
-we proceed by induction on F to prove ([[ F[FTop/x] ]]_v = [[ F ]]_v).
+suppose ([[ FAtom x ]]v = 1) (H).
+we proceed by induction on F to prove ([[ F[FTop/x] ]]v = [[ F ]]v).
case FBot.
- the thesis becomes ([[ FBot[FTop/x] ]]_v = [[ FBot ]]_v).
- the thesis becomes ([[ FBot ]]_v = [[ FBot ]]_v).
+ the thesis becomes ([[ FBot[FTop/x] ]]v = [[ FBot ]]v).
+ the thesis becomes ([[ FBot ]]v = [[ FBot ]]v).
done.
case FTop.
- the thesis becomes ([[ FTop[FTop/x] ]]_v = [[ FTop ]]_v).
- the thesis becomes ([[ FTop ]]_v = [[ FTop ]]_v).
+ the thesis becomes ([[ FTop[FTop/x] ]]v = [[ FTop ]]v).
+ the thesis becomes ([[ FTop ]]v = [[ FTop ]]v).
done.
case FAtom.
assume n : ℕ.
- the thesis becomes ([[ (FAtom n)[FTop/x] ]]_v = [[ FAtom n ]]_v).
- the thesis becomes ([[ if eqb n x then FTop else (FAtom n) ]]_v = [[ FAtom n ]]_v).
+ the thesis becomes ([[ (FAtom n)[FTop/x] ]]v = [[ FAtom n ]]v).
+ the thesis becomes ([[ if eqb n x then FTop else (FAtom n) ]]v = [[ FAtom n ]]v).
by decidable_eq_nat we proved (n = x ∨ n ≠ x) (H1).
- we proceed by cases on H1 to prove ([[ if eqb n x then FTop else (FAtom n) ]]_v = [[ FAtom n ]]_v).
+ we proceed by cases on H1 to prove ([[ if eqb n x then FTop else (FAtom n) ]]v = [[ FAtom n ]]v).
case Left.
by H2, eq_to_eqb_true we proved (eqb n x = true) (H3).
conclude
- ([[ if eqb n x then FTop else (FAtom n) ]]_v)
- = ([[ if true then FTop else (FAtom n) ]]_v) by H3.
- = ([[ FTop ]]_v).
+ ([[ if eqb n x then FTop else (FAtom n) ]]v)
+ = ([[ if true then FTop else (FAtom n) ]]v) by H3.
+ = ([[ FTop ]]v).
= 1.
- = ([[ FAtom x ]]_v) by H.
- = ([[ FAtom n ]]_v) by H2.
+ = ([[ FAtom x ]]v) by H.
+ = ([[ FAtom n ]]v) by H2.
done.
case Right.
by H2, not_eq_to_eqb_false we proved (eqb n x = false) (H3).
conclude
- ([[ if eqb n x then FTop else (FAtom n) ]]_v)
- = ([[ if false then FTop else (FAtom n) ]]_v) by H3.
- = ([[ FAtom n ]]_v).
+ ([[ if eqb n x then FTop else (FAtom n) ]]v)
+ = ([[ if false then FTop else (FAtom n) ]]v) by H3.
+ = ([[ FAtom n ]]v).
done.
case FAnd.
assume f1 : Formula.
- by induction hypothesis we know ([[ f1[FTop/x] ]]_v = [[ f1 ]]_v) (H1).
+ by induction hypothesis we know ([[ f1[FTop/x] ]]v = [[ f1 ]]v) (H1).
assume f2 : Formula.
- by induction hypothesis we know ([[ f2[FTop/x] ]]_v = [[ f2 ]]_v) (H2).
- the thesis becomes ([[ (FAnd f1 f2)[FTop/x] ]]_v = [[ FAnd f1 f2 ]]_v).
+ by induction hypothesis we know ([[ f2[FTop/x] ]]v = [[ f2 ]]v) (H2).
+ the thesis becomes ([[ (FAnd f1 f2)[FTop/x] ]]v = [[ FAnd f1 f2 ]]v).
conclude
- ([[ (FAnd f1 f2)[FTop/x] ]]_v)
- = ([[ FAnd (f1[FTop/x]) (f2[FTop/x]) ]]_v).
- = (min [[ f1[FTop/x] ]]_v [[ f2[FTop/x] ]]_v).
- = (min [[ f1 ]]_v [[ f2[FTop/x] ]]_v) by H1.
- = (min [[ f1 ]]_v [[ f2 ]]_v) by H2.
- = ([[ FAnd f1 f2 ]]_v).
+ ([[ (FAnd f1 f2)[FTop/x] ]]v)
+ = ([[ FAnd (f1[FTop/x]) (f2[FTop/x]) ]]v).
+ = (min [[ f1[FTop/x] ]]v [[ f2[FTop/x] ]]v).
+ = (min [[ f1 ]]v [[ f2[FTop/x] ]]v) by H1.
+ = (min [[ f1 ]]v [[ f2 ]]v) by H2.
+ = ([[ FAnd f1 f2 ]]v).
done.
case FOr.
assume f1 : Formula.
- by induction hypothesis we know ([[ f1[FTop/x] ]]_v = [[ f1 ]]_v) (H1).
+ by induction hypothesis we know ([[ f1[FTop/x] ]]v = [[ f1 ]]v) (H1).
assume f2 : Formula.
- by induction hypothesis we know ([[ f2[FTop/x] ]]_v = [[ f2 ]]_v) (H2).
- the thesis becomes ([[ (FOr f1 f2)[FTop/x] ]]_v = [[ FOr f1 f2 ]]_v).
+ by induction hypothesis we know ([[ f2[FTop/x] ]]v = [[ f2 ]]v) (H2).
+ the thesis becomes ([[ (FOr f1 f2)[FTop/x] ]]v = [[ FOr f1 f2 ]]v).
conclude
- ([[ (FOr f1 f2)[FTop/x] ]]_v)
- = ([[ FOr (f1[FTop/x]) (f2[FTop/x]) ]]_v).
- = (max [[ f1[FTop/x] ]]_v [[ f2[FTop/x] ]]_v).
- = (max [[ f1 ]]_v [[ f2[FTop/x] ]]_v) by H1.
- = (max [[ f1 ]]_v [[ f2 ]]_v) by H2.
- = ([[ FOr f1 f2 ]]_v).
+ ([[ (FOr f1 f2)[FTop/x] ]]v)
+ = ([[ FOr (f1[FTop/x]) (f2[FTop/x]) ]]v).
+ = (max [[ f1[FTop/x] ]]v [[ f2[FTop/x] ]]v).
+ = (max [[ f1 ]]v [[ f2[FTop/x] ]]v) by H1.
+ = (max [[ f1 ]]v [[ f2 ]]v) by H2.
+ = ([[ FOr f1 f2 ]]v).
done.
case FImpl.
assume f1 : Formula.
- by induction hypothesis we know ([[ f1[FTop/x] ]]_v = [[ f1 ]]_v) (H1).
+ by induction hypothesis we know ([[ f1[FTop/x] ]]v = [[ f1 ]]v) (H1).
assume f2 : Formula.
- by induction hypothesis we know ([[ f2[FTop/x] ]]_v = [[ f2 ]]_v) (H2).
- the thesis becomes ([[ (FImpl f1 f2)[FTop/x] ]]_v = [[ FImpl f1 f2 ]]_v).
+ by induction hypothesis we know ([[ f2[FTop/x] ]]v = [[ f2 ]]v) (H2).
+ the thesis becomes ([[ (FImpl f1 f2)[FTop/x] ]]v = [[ FImpl f1 f2 ]]v).
conclude
- ([[ (FImpl f1 f2)[FTop/x] ]]_v)
- = ([[ FImpl (f1[FTop/x]) (f2[FTop/x]) ]]_v).
- = (max (1 - [[ f1[FTop/x] ]]_v) [[ f2[FTop/x] ]]_v).
- = (max (1 - [[ f1 ]]_v) [[ f2[FTop/x] ]]_v) by H1.
- = (max (1 - [[ f1 ]]_v) [[ f2 ]]_v) by H2.
- = ([[ FImpl f1 f2 ]]_v).
+ ([[ (FImpl f1 f2)[FTop/x] ]]v)
+ = ([[ FImpl (f1[FTop/x]) (f2[FTop/x]) ]]v).
+ = (max (1 - [[ f1[FTop/x] ]]v) [[ f2[FTop/x] ]]v).
+ = (max (1 - [[ f1 ]]v) [[ f2[FTop/x] ]]v) by H1.
+ = (max (1 - [[ f1 ]]v) [[ f2 ]]v) by H2.
+ = ([[ FImpl f1 f2 ]]v).
done.
case FNot.
assume f : Formula.
- by induction hypothesis we know ([[ f[FTop/x] ]]_v = [[ f ]]_v) (H1).
- the thesis becomes ([[ (FNot f)[FTop/x] ]]_v = [[ FNot f ]]_v).
+ by induction hypothesis we know ([[ f[FTop/x] ]]v = [[ f ]]v) (H1).
+ the thesis becomes ([[ (FNot f)[FTop/x] ]]v = [[ FNot f ]]v).
conclude
- ([[ (FNot f)[FTop/x] ]]_v)
- = ([[ FNot (f[FTop/x]) ]]_v).
- = (1 - [[ f[FTop/x] ]]_v).
- = (1 - [[ f ]]_v) by H1.
- = ([[ FNot f ]]_v).
+ ([[ (FNot f)[FTop/x] ]]v)
+ = ([[ FNot (f[FTop/x]) ]]v).
+ = (1 - [[ f[FTop/x] ]]v).
+ = (1 - [[ f ]]v) by H1.
+ = ([[ FNot f ]]v).
done.
(*END*)
qed.
assume F : Formula.
assume x : ℕ.
assume v : (ℕ → ℕ).
-the thesis becomes ([[ IFTE (FAtom x) (F[FTop/x]) (F[FBot/x])]]_v = [[ F ]]_v).
-by sem_bool we proved ([[ FAtom x]]_v = 0 ∨ [[ FAtom x]]_v = 1) (H).
-we proceed by cases on H to prove ([[ IFTE (FAtom x) (F[FTop/x]) (F[FBot/x])]]_v = [[ F ]]_v).
+the thesis becomes ([[ IFTE (FAtom x) (F[FTop/x]) (F[FBot/x])]]v = [[ F ]]v).
+by sem_bool we proved ([[ FAtom x]]v = 0 ∨ [[ FAtom x]]v = 1) (H).
+we proceed by cases on H to prove ([[ IFTE (FAtom x) (F[FTop/x]) (F[FBot/x])]]v = [[ F ]]v).
case Left.
conclude
- ([[ IFTE (FAtom x) (F[FTop/x]) (F[FBot/x])]]_v)
- = ([[ FOr (FAnd (FAtom x) (F[FTop/x])) (FAnd (FNot (FAtom x)) (F[FBot/x]))]]_v).
- = (max [[ (FAnd (FAtom x) (F[FTop/x])) ]]_v [[ (FAnd (FNot (FAtom x)) (F[FBot/x]))]]_v).
- = (max (min [[ FAtom x ]]_v [[ F[FTop/x] ]]_v) (min (1 - [[ FAtom x ]]_v) [[ F[FBot/x] ]]_v)).
- = (max (min 0 [[ F[FTop/x] ]]_v) (min (1 - 0) [[ F[FBot/x] ]]_v)) by H1.
- = (max 0 (min 1 [[ F[FBot/x] ]]_v)).
- = (max 0 [[ F[FBot/x] ]]_v) by min_1_sem.
- = ([[ F[FBot/x] ]]_v).
- = ([[ F ]]_v) by H1, shannon_false.
+ ([[ IFTE (FAtom x) (F[FTop/x]) (F[FBot/x])]]v)
+ = ([[ FOr (FAnd (FAtom x) (F[FTop/x])) (FAnd (FNot (FAtom x)) (F[FBot/x]))]]v).
+ = (max [[ (FAnd (FAtom x) (F[FTop/x])) ]]v [[ (FAnd (FNot (FAtom x)) (F[FBot/x]))]]v).
+ = (max (min [[ FAtom x ]]v [[ F[FTop/x] ]]v) (min (1 - [[ FAtom x ]]v) [[ F[FBot/x] ]]v)).
+ = (max (min 0 [[ F[FTop/x] ]]v) (min (1 - 0) [[ F[FBot/x] ]]v)) by H1.
+ = (max 0 (min 1 [[ F[FBot/x] ]]v)).
+ = (max 0 [[ F[FBot/x] ]]v) by min_1_sem.
+ = ([[ F[FBot/x] ]]v).
+ = ([[ F ]]v) by H1, shannon_false.
done.
case Right.
conclude
- ([[ IFTE (FAtom x) (F[FTop/x]) (F[FBot/x])]]_v)
- = ([[ FOr (FAnd (FAtom x) (F[FTop/x])) (FAnd (FNot (FAtom x)) (F[FBot/x]))]]_v).
- = (max [[ (FAnd (FAtom x) (F[FTop/x])) ]]_v [[ (FAnd (FNot (FAtom x)) (F[FBot/x]))]]_v).
- = (max (min [[ FAtom x ]]_v [[ F[FTop/x] ]]_v) (min (1 - [[ FAtom x ]]_v) [[ F[FBot/x] ]]_v)).
- = (max (min 1 [[ F[FTop/x] ]]_v) (min (1 - 1) [[ F[FBot/x] ]]_v)) by H1.
- = (max (min 1 [[ F[FTop/x] ]]_v) (min 0 [[ F[FBot/x] ]]_v)).
- = (max [[ F[FTop/x] ]]_v (min 0 [[ F[FBot/x] ]]_v)) by min_1_sem.
- = (max [[ F[FTop/x] ]]_v 0).
- = ([[ F[FTop/x] ]]_v) by max_0_sem.
- = ([[ F ]]_v) by H1, shannon_true.
+ ([[ IFTE (FAtom x) (F[FTop/x]) (F[FBot/x])]]v)
+ = ([[ FOr (FAnd (FAtom x) (F[FTop/x])) (FAnd (FNot (FAtom x)) (F[FBot/x]))]]v).
+ = (max [[ (FAnd (FAtom x) (F[FTop/x])) ]]v [[ (FAnd (FNot (FAtom x)) (F[FBot/x]))]]v).
+ = (max (min [[ FAtom x ]]v [[ F[FTop/x] ]]v) (min (1 - [[ FAtom x ]]v) [[ F[FBot/x] ]]v)).
+ = (max (min 1 [[ F[FTop/x] ]]v) (min (1 - 1) [[ F[FBot/x] ]]v)) by H1.
+ = (max (min 1 [[ F[FTop/x] ]]v) (min 0 [[ F[FBot/x] ]]v)).
+ = (max [[ F[FTop/x] ]]v (min 0 [[ F[FBot/x] ]]v)) by min_1_sem.
+ = (max [[ F[FTop/x] ]]v 0).
+ = ([[ F[FTop/x] ]]v) by max_0_sem.
+ = ([[ F ]]v) by H1, shannon_true.
done.
(*END*)
qed.
Il comando conclude lavora SOLO sulla parte sinistra della tesi. È il comando
con cui si inizia una catena di uguaglianze. La prima formula che si
scrive deve essere esattamente uguale alla parte sinistra della conclusione
- originale. Esempio `conclude ([[ FAtom x ]]_v) = ([[ FAtom n ]]_v) by H.`
+ originale. Esempio `conclude ([[ FAtom x ]]v) = ([[ FAtom n ]]v) by H.`
Se la giustificazione non è un lemma o una ipotesi ma la semplice espansione
di una definizione, la parte `by ...` deve essere omessa.
Un esempio di espressione è `eqb n m`, che confronta i due numeri naturali
`n` ed `m`.
- * [[ formula ]]_v
+ * [[ formula ]]v
Questa notazione utilizza la funzione `sem` precedentemente definita, in
- particolare `[[ f ]]_v` è una abbreviazione per `sem v f`.
+ particolare `[[ f ]]v` è una abbreviazione per `sem v f`.
ATTENZIONE
definition if_then_else ≝ λT:Type.λe,t,f.match e return λ_.T with [ true ⇒ t | false ⇒ f].
notation > "'if' term 19 e 'then' term 19 t 'else' term 90 f" non associative with precedence 90 for @{ 'if_then_else $e $t $f }.
notation < "'if' \nbsp term 19 e \nbsp 'then' \nbsp term 19 t \nbsp 'else' \nbsp term 90 f \nbsp" non associative with precedence 90 for @{ 'if_then_else $e $t $f }.
-interpretation "Formula if_then_else" 'if_then_else e t f = (if_then_else _ e t f).
+interpretation "Formula if_then_else" 'if_then_else e t f = (if_then_else ? e t f).
notation < "[[ \nbsp term 19 a \nbsp ]] \nbsp \sub term 90 v" non associative with precedence 90 for @{ 'semantics $v $a }.
notation > "[[ term 19 a ]] \sub term 90 v" non associative with precedence 90 for @{ 'semantics $v $a }.
-notation > "[[ term 19 a ]]_ term 90 v" non associative with precedence 90 for @{ sem $v $a }.
+notation > "[[ term 19 a ]] term 90 v" non associative with precedence 90 for @{ sem $v $a }.
interpretation "Semantic of Formula" 'semantics v a = (sem v a).
Tale formula è valida per la funzione di valutazione `v1101`.
- Il comando `eval normalize [[ esempio1 ]]_v1101` permette di calcolare
+ Il comando `eval normalize [[ esempio1 ]]v1101` permette di calcolare
la funzione `sem` che avete appena definito. Tale funzione deve
computare a 1 (verrà mostrata una finestra col risultato).
Se così non fosse significa che avete commesso un errore nella
(* Decommenta ed esegui. *)
-(* eval normalize on [[ esempio1 ]]_v1101. *)
+(* eval normalize on [[ esempio1 ]]v1101. *)
(* Esercizio 3
* F ≡ G
- Questa notazione è una abbreviazione per `∀v.[[ f ]]_v = [[ g ]]_v`.
+ Questa notazione è una abbreviazione per `∀v.[[ f ]]v = [[ g ]]v`.
Asserisce che for ogni funzione di valutazione `v`, la semantica di `f`
in `v` è uguale alla semantica di `g` in `v`.
notation < "t [ \nbsp term 19 a / term 19 b \nbsp ]" non associative with precedence 90 for @{ 'substitution $b $a $t }.
notation > "t [ term 90 a / term 90 b]" non associative with precedence 90 for @{ 'substitution $b $a $t }.
interpretation "Substitution for Formula" 'substitution b a t = (subst b a t).
-definition equiv ≝ λF1,F2. ∀v.[[ F1 ]]_v = [[ F2 ]]_v.
+definition equiv ≝ λF1,F2. ∀v.[[ F1 ]]v = [[ F2 ]]v.
notation "hvbox(a \nbsp break mstyle color #0000ff (≡) \nbsp b)" non associative with precedence 45 for @{ 'equivF $a $b }.
notation > "a ≡ b" non associative with precedence 50 for @{ equiv $a $b }.
interpretation "equivalence for Formulas" 'equivF a b = (equiv a b).
the thesis becomes (FBot[ G1/x ] ≡ FBot[ G2/x ]).
the thesis becomes (FBot ≡ FBot[ G2/x ]).
the thesis becomes (FBot ≡ FBot).
- the thesis becomes (∀v.[[FBot]]_v = [[FBot]]_v).
+ the thesis becomes (∀v.[[FBot]]v = [[FBot]]v).
assume v : (ℕ → ℕ).
- the thesis becomes (0 = [[FBot]]_v).
+ the thesis becomes (0 = [[FBot]]v).
the thesis becomes (0 = 0).
done.
case FTop.
(*BEGIN*)
the thesis becomes (FTop[ G1/x ] ≡ FTop[ G2/x ]).
the thesis becomes (FTop ≡ FTop).
- the thesis becomes (∀v. [[FTop]]_v = [[FTop]]_v).
+ the thesis becomes (∀v. [[FTop]]v = [[FTop]]v).
assume v : (ℕ → ℕ).
the thesis becomes (1 = 1).
(*END*)
case false.
(*BEGIN*)
the thesis becomes (FAtom n ≡ FAtom n).
- the thesis becomes (∀v. [[FAtom n]]_v = [[FAtom n]]_v).
+ the thesis becomes (∀v. [[FAtom n]]v = [[FAtom n]]v).
assume v : (ℕ → ℕ).
the thesis becomes (v n = v n).
(*END*)
assume F2 : Formula.
by induction hypothesis we know (F2[ G1/x ] ≡ F2[ G2/x ]) (IH2).
the thesis becomes
- (∀v.[[ (FAnd F1 F2)[ G1/x ] ]]_v = [[ (FAnd F1 F2)[ G2/x ] ]]_v).
+ (∀v.[[ (FAnd F1 F2)[ G1/x ] ]]v = [[ (FAnd F1 F2)[ G2/x ] ]]v).
assume v : (ℕ → ℕ).
the thesis becomes
- (min ([[ F1[ G1/x ] ]]_v) ([[ F2[ G1/x ] ]]_v) =
- min ([[ F1[ G2/x ] ]]_v) ([[ F2[ G2/x ] ]]_v)).
- by IH1 we proved (∀v1.[[ F1[ G1/x ] ]]_v1 = [[ F1[ G2/x ] ]]_v1) (IH11).
- by (*BEGIN*)IH2(*END*) we proved (∀v2.[[ F2[ G1/x ] ]]_v2 = [[ F2[ G2/x ] ]]_v2) (IH22).
- by IH11 we proved ([[ F1[ G1/x ] ]]_v = [[ F1[ G2/x ] ]]_v) (IH111).
- by (*BEGIN*)IH22(*END*) we proved ([[ F2[ G1/x ] ]]_v = [[ F2[ G2/x ] ]]_v) (IH222).
+ (min ([[ F1[ G1/x ] ]]v) ([[ F2[ G1/x ] ]]v) =
+ min ([[ F1[ G2/x ] ]]v) ([[ F2[ G2/x ] ]]v)).
+ by IH1 we proved (∀v1.[[ F1[ G1/x ] ]]v1 = [[ F1[ G2/x ] ]]v1) (IH11).
+ by (*BEGIN*)IH2(*END*) we proved (∀v2.[[ F2[ G1/x ] ]]v2 = [[ F2[ G2/x ] ]]v2) (IH22).
+ by IH11 we proved ([[ F1[ G1/x ] ]]v = [[ F1[ G2/x ] ]]v) (IH111).
+ by (*BEGIN*)IH22(*END*) we proved ([[ F2[ G1/x ] ]]v = [[ F2[ G2/x ] ]]v) (IH222).
conclude
- (min ([[ F1[ G1/x ] ]]_v) ([[ F2[ G1/x ] ]]_v))
- = (min ([[ F1[ G2/x ] ]]_v) ([[ F2[ G1/x ] ]]_v)) by IH111.
- = (min ([[(F1[ G2/x ])]]_v) ([[(F2[ G2/x ])]]_v)) by (*BEGIN*)IH222(*END*).
+ (min ([[ F1[ G1/x ] ]]v) ([[ F2[ G1/x ] ]]v))
+ = (min ([[ F1[ G2/x ] ]]v) ([[ F2[ G1/x ] ]]v)) by IH111.
+ = (min ([[(F1[ G2/x ])]]v) ([[(F2[ G2/x ])]]v)) by (*BEGIN*)IH222(*END*).
(*END*)
done.
case FOr.
assume F2 : Formula.
by induction hypothesis we know (F2[ G1/x ] ≡ F2[ G2/x ]) (IH2).
the thesis becomes
- (∀v.[[ (FOr F1 F2)[ G1/x ] ]]_v = [[ (FOr F1 F2)[ G2/x ] ]]_v).
+ (∀v.[[ (FOr F1 F2)[ G1/x ] ]]v = [[ (FOr F1 F2)[ G2/x ] ]]v).
assume v : (ℕ → ℕ).
the thesis becomes
- (max ([[ F1[ G1/x ] ]]_v) ([[ F2[ G1/x ] ]]_v) =
- max ([[ F1[ G2/x ] ]]_v) ([[ F2[ G2/x ] ]]_v)).
- by IH1 we proved (∀v1.[[ F1[ G1/x ] ]]_v1 = [[ F1[ G2/x ] ]]_v1) (IH11).
- by IH2 we proved (∀v2.[[ F2[ G1/x ] ]]_v2 = [[ F2[ G2/x ] ]]_v2) (IH22).
- by IH11 we proved ([[ F1[ G1/x ] ]]_v = [[ F1[ G2/x ] ]]_v) (IH111).
- by IH22 we proved ([[ F2[ G1/x ] ]]_v = [[ F2[ G2/x ] ]]_v) (IH222).
+ (max ([[ F1[ G1/x ] ]]v) ([[ F2[ G1/x ] ]]v) =
+ max ([[ F1[ G2/x ] ]]v) ([[ F2[ G2/x ] ]]v)).
+ by IH1 we proved (∀v1.[[ F1[ G1/x ] ]]v1 = [[ F1[ G2/x ] ]]v1) (IH11).
+ by IH2 we proved (∀v2.[[ F2[ G1/x ] ]]v2 = [[ F2[ G2/x ] ]]v2) (IH22).
+ by IH11 we proved ([[ F1[ G1/x ] ]]v = [[ F1[ G2/x ] ]]v) (IH111).
+ by IH22 we proved ([[ F2[ G1/x ] ]]v = [[ F2[ G2/x ] ]]v) (IH222).
conclude
- (max ([[ F1[ G1/x ] ]]_v) ([[ F2[ G1/x ] ]]_v))
- = (max ([[ F1[ G2/x ] ]]_v) ([[ F2[ G1/x ] ]]_v)) by IH111.
- = (max ([[(F1[ G2/x ])]]_v) ([[(F2[ G2/x ])]]_v)) by IH222.
+ (max ([[ F1[ G1/x ] ]]v) ([[ F2[ G1/x ] ]]v))
+ = (max ([[ F1[ G2/x ] ]]v) ([[ F2[ G1/x ] ]]v)) by IH111.
+ = (max ([[(F1[ G2/x ])]]v) ([[(F2[ G2/x ])]]v)) by IH222.
(*END*)
done.
case FImpl.
assume F2 : Formula.
by induction hypothesis we know (F2[ G1/x ] ≡ F2[ G2/x ]) (IH2).
the thesis becomes
- (∀v.max (1 - [[ F1[ G1/x ] ]]_v) ([[ F2[ G1/x ] ]]_v) =
- max (1 - [[ F1[ G2/x ] ]]_v) ([[ F2[ G2/x ] ]]_v)).
+ (∀v.max (1 - [[ F1[ G1/x ] ]]v) ([[ F2[ G1/x ] ]]v) =
+ max (1 - [[ F1[ G2/x ] ]]v) ([[ F2[ G2/x ] ]]v)).
assume v : (ℕ → ℕ).
- by IH1 we proved ([[ F1[ G1/x ] ]]_v = [[ F1[ G2/x ] ]]_v) (IH11).
- by IH2 we proved ([[ F2[ G1/x ] ]]_v = [[ F2[ G2/x ] ]]_v) (IH22).
+ by IH1 we proved ([[ F1[ G1/x ] ]]v = [[ F1[ G2/x ] ]]v) (IH11).
+ by IH2 we proved ([[ F2[ G1/x ] ]]v = [[ F2[ G2/x ] ]]v) (IH22).
conclude
- (max (1-[[ F1[ G1/x ] ]]_v) ([[ F2[ G1/x ] ]]_v))
- = (max (1-[[ F1[ G2/x ] ]]_v) ([[ F2[ G1/x ] ]]_v)) by IH11.
- = (max (1-[[ F1[ G2/x ] ]]_v) ([[ F2[ G2/x ] ]]_v)) by IH22.
+ (max (1-[[ F1[ G1/x ] ]]v) ([[ F2[ G1/x ] ]]v))
+ = (max (1-[[ F1[ G2/x ] ]]v) ([[ F2[ G1/x ] ]]v)) by IH11.
+ = (max (1-[[ F1[ G2/x ] ]]v) ([[ F2[ G2/x ] ]]v)) by IH22.
done.
case FNot.
(*BEGIN*)
assume F1 : Formula.
by induction hypothesis we know (F1[ G1/x ] ≡ F1[ G2/x ]) (IH).
the thesis becomes (FNot (F1[ G1/x ]) ≡ FNot (F1[ G2/x ])).
- the thesis becomes (∀v.[[FNot (F1[ G1/x ])]]_v = [[FNot (F1[ G2/x ])]]_v).
+ the thesis becomes (∀v.[[FNot (F1[ G1/x ])]]v = [[FNot (F1[ G2/x ])]]v).
assume v : (ℕ → ℕ).
- the thesis becomes (1 - [[F1[ G1/x ]]]_v = [[FNot (F1[ G2/x ])]]_v).
- the thesis becomes (1 - [[ F1[ G1/x ] ]]_v = 1 - [[ F1[ G2/x ] ]]_v).
- by IH we proved (∀v1.[[ F1[ G1/x ] ]]_v1 = [[ F1[ G2/x ] ]]_v1) (IH1).
- by IH1 we proved ([[ F1[ G1/x ] ]]_v = [[ F1[ G2/x ] ]]_v) (IH2).
+ the thesis becomes (1 - [[F1[ G1/x ]]]v = [[FNot (F1[ G2/x ])]]v).
+ the thesis becomes (1 - [[ F1[ G1/x ] ]]v = 1 - [[ F1[ G2/x ] ]]v).
+ by IH we proved (∀v1.[[ F1[ G1/x ] ]]v1 = [[ F1[ G2/x ] ]]v1) (IH1).
+ by IH1 we proved ([[ F1[ G1/x ] ]]v = [[ F1[ G2/x ] ]]v) (IH2).
conclude
- (1-[[ F1[ G1/x ] ]]_v)
- = (1-[[ F1[ G2/x ] ]]_v) by IH2.
+ (1-[[ F1[ G1/x ] ]]v)
+ = (1-[[ F1[ G2/x ] ]]v) by IH2.
(*END*)
done.
qed.
for @{ 'Forall (λ${ident x}.$Px) }.
notation < "\forall ident x.break term 19 Px" with precedence 20
for @{ 'Forall (λ${ident x}:$tx.$Px) }.
-interpretation "Forall" 'Forall \eta.Px = (Forall _ Px).
+interpretation "Forall" 'Forall \eta.Px = (Forall ? Px).
(* Variables *)
axiom A : CProp.
(* begin a proof: draws the root *)
notation > "'prove' p" non associative with precedence 19
for @{ 'prove $p }.
-interpretation "prove KO" 'prove p = (cast _ _ (show p _)).
-interpretation "prove OK" 'prove p = (show p _).
+interpretation "prove KO" 'prove p = (cast ? ? (show p ?)).
+interpretation "prove OK" 'prove p = (show p ?).
(* Leaves *)
notation < "\infrule (t\atop ⋮) a ?" with precedence 19
interpretation "leaf OK" 'leaf_ok a t = (show a t).
notation < "\infrule (t\atop ⋮) mstyle color #ff0000 (a) ?" with precedence 19
for @{ 'leaf_ko $a $t }.
-interpretation "leaf KO" 'leaf_ko a t = (cast _ _ (show a t)).
+interpretation "leaf KO" 'leaf_ko a t = (cast ? ? (show a t)).
(* discharging *)
notation < "[ a ] \sup mstyle color #ff0000 (H)" with precedence 19
for @{ 'discharge_ko_1 $a $H }.
interpretation "discharge_ko_1" 'discharge_ko_1 a H =
- (show a (cast _ _ (Discharge _ H))).
+ (show a (cast ? ? (Discharge ? H))).
notation < "[ mstyle color #ff0000 (a) ] \sup mstyle color #ff0000 (H)" with precedence 19
for @{ 'discharge_ko_2 $a $H }.
interpretation "discharge_ko_2" 'discharge_ko_2 a H =
- (cast _ _ (show a (cast _ _ (Discharge _ H)))).
+ (cast ? ? (show a (cast ? ? (Discharge ? H)))).
notation < "[ a ] \sup H" with precedence 19
for @{ 'discharge_ok_1 $a $H }.
interpretation "discharge_ok_1" 'discharge_ok_1 a H =
- (show a (Discharge _ H)).
+ (show a (Discharge ? H)).
notation < "[ mstyle color #ff0000 (a) ] \sup H" with precedence 19
for @{ 'discharge_ok_2 $a $H }.
interpretation "discharge_ok_2" 'discharge_ok_2 a H =
- (cast _ _ (show a (Discharge _ H))).
+ (cast ? ? (show a (Discharge ? H))).
notation > "'discharge' [H]" with precedence 19
for @{ 'discharge $H }.
-interpretation "discharge KO" 'discharge H = (cast _ _ (Discharge _ H)).
-interpretation "discharge OK" 'discharge H = (Discharge _ H).
+interpretation "discharge KO" 'discharge H = (cast ? ? (Discharge ? H)).
+interpretation "discharge OK" 'discharge H = (Discharge ? H).
(* ⇒ introduction *)
notation < "\infrule hbox(\emsp b \emsp) ab (mstyle color #ff0000 (⇒\sub\i \emsp) ident H) " with precedence 19
for @{ 'Imply_intro_ko_1 $ab (λ${ident H}:$p.$b) }.
interpretation "Imply_intro_ko_1" 'Imply_intro_ko_1 ab \eta.b =
- (show ab (cast _ _ (Imply_intro _ _ b))).
+ (show ab (cast ? ? (Imply_intro ? ? b))).
notation < "\infrule hbox(\emsp b \emsp) mstyle color #ff0000 (ab) (mstyle color #ff0000 (⇒\sub\i \emsp) ident H) " with precedence 19
for @{ 'Imply_intro_ko_2 $ab (λ${ident H}:$p.$b) }.
interpretation "Imply_intro_ko_2" 'Imply_intro_ko_2 ab \eta.b =
- (cast _ _ (show ab (cast _ _ (Imply_intro _ _ b)))).
+ (cast ? ? (show ab (cast ? ? (Imply_intro ? ? b)))).
notation < "maction (\infrule hbox(\emsp b \emsp) ab (⇒\sub\i \emsp ident H) ) (\vdots)" with precedence 19
for @{ 'Imply_intro_ok_1 $ab (λ${ident H}:$p.$b) }.
interpretation "Imply_intro_ok_1" 'Imply_intro_ok_1 ab \eta.b =
- (show ab (Imply_intro _ _ b)).
+ (show ab (Imply_intro ? ? b)).
notation < "\infrule hbox(\emsp b \emsp) mstyle color #ff0000 (ab) (⇒\sub\i \emsp ident H) " with precedence 19
for @{ 'Imply_intro_ok_2 $ab (λ${ident H}:$p.$b) }.
interpretation "Imply_intro_ok_2" 'Imply_intro_ok_2 ab \eta.b =
- (cast _ _ (show ab (Imply_intro _ _ b))).
+ (cast ? ? (show ab (Imply_intro ? ? b))).
-notation > "⇒_'i' [ident H] term 90 b" with precedence 19
+notation > "⇒#'i' [ident H] term 90 b" with precedence 19
for @{ 'Imply_intro $b (λ${ident H}.show $b ?) }.
+
interpretation "Imply_intro KO" 'Imply_intro b pb =
- (cast _ (Imply unit b) (Imply_intro _ b pb)).
+ (cast ? (Imply unit b) (Imply_intro ? b pb)).
interpretation "Imply_intro OK" 'Imply_intro b pb =
- (Imply_intro _ b pb).
+ (Imply_intro ? b pb).
(* ⇒ elimination *)
notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) b mstyle color #ff0000 (⇒\sub\e) " with precedence 19
for @{ 'Imply_elim_ko_1 $ab $a $b }.
interpretation "Imply_elim_ko_1" 'Imply_elim_ko_1 ab a b =
- (show b (cast _ _ (Imply_elim _ _ (cast _ _ ab) (cast _ _ a)))).
+ (show b (cast ? ? (Imply_elim ? ? (cast ? ? ab) (cast ? ? a)))).
notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) mstyle color #ff0000 (b) mstyle color #ff0000 (⇒\sub\e) " with precedence 19
for @{ 'Imply_elim_ko_2 $ab $a $b }.
interpretation "Imply_elim_ko_2" 'Imply_elim_ko_2 ab a b =
- (cast _ _ (show b (cast _ _ (Imply_elim _ _ (cast _ _ ab) (cast _ _ a))))).
+ (cast ? ? (show b (cast ? ? (Imply_elim ? ? (cast ? ? ab) (cast ? ? a))))).
notation < "maction (\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) b (⇒\sub\e) ) (\vdots)" with precedence 19
for @{ 'Imply_elim_ok_1 $ab $a $b }.
interpretation "Imply_elim_ok_1" 'Imply_elim_ok_1 ab a b =
- (show b (Imply_elim _ _ ab a)).
+ (show b (Imply_elim ? ? ab a)).
notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) mstyle color #ff0000 (b) (⇒\sub\e) " with precedence 19
for @{ 'Imply_elim_ok_2 $ab $a $b }.
interpretation "Imply_elim_ok_2" 'Imply_elim_ok_2 ab a b =
- (cast _ _ (show b (Imply_elim _ _ ab a))).
+ (cast ? ? (show b (Imply_elim ? ? ab a))).
-notation > "⇒_'e' term 90 ab term 90 a" with precedence 19
+notation > "⇒#'e' term 90 ab term 90 a" with precedence 19
for @{ 'Imply_elim (show $ab ?) (show $a ?) }.
interpretation "Imply_elim KO" 'Imply_elim ab a =
- (cast _ _ (Imply_elim _ _ (cast (Imply unit unit) _ ab) (cast unit _ a))).
+ (cast ? ? (Imply_elim ? ? (cast (Imply unit unit) ? ab) (cast unit ? a))).
interpretation "Imply_elim OK" 'Imply_elim ab a =
- (Imply_elim _ _ ab a).
+ (Imply_elim ? ? ab a).
(* ∧ introduction *)
notation < "\infrule hbox(\emsp a \emsp\emsp\emsp b \emsp) ab mstyle color #ff0000 (∧\sub\i)" with precedence 19
for @{ 'And_intro_ko_1 $a $b $ab }.
interpretation "And_intro_ko_1" 'And_intro_ko_1 a b ab =
- (show ab (cast _ _ (And_intro _ _ a b))).
+ (show ab (cast ? ? (And_intro ? ? a b))).
notation < "\infrule hbox(\emsp a \emsp\emsp\emsp b \emsp) mstyle color #ff0000 (ab) mstyle color #ff0000 (∧\sub\i)" with precedence 19
for @{ 'And_intro_ko_2 $a $b $ab }.
interpretation "And_intro_ko_2" 'And_intro_ko_2 a b ab =
- (cast _ _ (show ab (cast _ _ (And_intro _ _ a b)))).
+ (cast ? ? (show ab (cast ? ? (And_intro ? ? a b)))).
notation < "maction (\infrule hbox(\emsp a \emsp\emsp\emsp b \emsp) ab (∧\sub\i)) (\vdots)" with precedence 19
for @{ 'And_intro_ok_1 $a $b $ab }.
interpretation "And_intro_ok_1" 'And_intro_ok_1 a b ab =
- (show ab (And_intro _ _ a b)).
+ (show ab (And_intro ? ? a b)).
notation < "\infrule hbox(\emsp a \emsp\emsp\emsp b \emsp) mstyle color #ff0000 (ab) (∧\sub\i)" with precedence 19
for @{ 'And_intro_ok_2 $a $b $ab }.
interpretation "And_intro_ok_2" 'And_intro_ok_2 a b ab =
- (cast _ _ (show ab (And_intro _ _ a b))).
+ (cast ? ? (show ab (And_intro ? ? a b))).
-notation > "∧_'i' term 90 a term 90 b" with precedence 19
+notation > "∧#'i' term 90 a term 90 b" with precedence 19
for @{ 'And_intro (show $a ?) (show $b ?) }.
-interpretation "And_intro KO" 'And_intro a b = (cast _ _ (And_intro _ _ a b)).
-interpretation "And_intro OK" 'And_intro a b = (And_intro _ _ a b).
+interpretation "And_intro KO" 'And_intro a b = (cast ? ? (And_intro ? ? a b)).
+interpretation "And_intro OK" 'And_intro a b = (And_intro ? ? a b).
(* ∧ elimination *)
notation < "\infrule hbox(\emsp ab \emsp) a mstyle color #ff0000 (∧\sub(\e_\l))" with precedence 19
for @{ 'And_elim_l_ko_1 $ab $a }.
interpretation "And_elim_l_ko_1" 'And_elim_l_ko_1 ab a =
- (show a (cast _ _ (And_elim_l _ _ (cast _ _ ab)))).
+ (show a (cast ? ? (And_elim_l ? ? (cast ? ? ab)))).
notation < "\infrule hbox(\emsp ab \emsp) mstyle color #ff0000 (a) mstyle color #ff0000 (∧\sub(\e_\l))" with precedence 19
for @{ 'And_elim_l_ko_2 $ab $a }.
interpretation "And_elim_l_ko_2" 'And_elim_l_ko_2 ab a =
- (cast _ _ (show a (cast _ _ (And_elim_l _ _ (cast _ _ ab))))).
+ (cast ? ? (show a (cast ? ? (And_elim_l ? ? (cast ? ? ab))))).
notation < "maction (\infrule hbox(\emsp ab \emsp) a (∧\sub(\e_\l))) (\vdots)" with precedence 19
for @{ 'And_elim_l_ok_1 $ab $a }.
interpretation "And_elim_l_ok_1" 'And_elim_l_ok_1 ab a =
- (show a (And_elim_l _ _ ab)).
+ (show a (And_elim_l ? ? ab)).
notation < "\infrule hbox(\emsp ab \emsp) mstyle color #ff0000 (a) (∧\sub(\e_\l))" with precedence 19
for @{ 'And_elim_l_ok_2 $ab $a }.
interpretation "And_elim_l_ok_2" 'And_elim_l_ok_2 ab a =
- (cast _ _ (show a (And_elim_l _ _ ab))).
+ (cast ? ? (show a (And_elim_l ? ? ab))).
-notation > "∧_'e_l' term 90 ab" with precedence 19
+notation > "∧#'e_l' term 90 ab" with precedence 19
for @{ 'And_elim_l (show $ab ?) }.
-interpretation "And_elim_l KO" 'And_elim_l a = (cast _ _ (And_elim_l _ _ (cast (And unit unit) _ a))).
-interpretation "And_elim_l OK" 'And_elim_l a = (And_elim_l _ _ a).
+interpretation "And_elim_l KO" 'And_elim_l a = (cast ? ? (And_elim_l ? ? (cast (And unit unit) ? a))).
+interpretation "And_elim_l OK" 'And_elim_l a = (And_elim_l ? ? a).
notation < "\infrule hbox(\emsp ab \emsp) a mstyle color #ff0000 (∧\sub(\e_\r))" with precedence 19
for @{ 'And_elim_r_ko_1 $ab $a }.
interpretation "And_elim_r_ko_1" 'And_elim_r_ko_1 ab a =
- (show a (cast _ _ (And_elim_r _ _ (cast _ _ ab)))).
+ (show a (cast ? ? (And_elim_r ? ? (cast ? ? ab)))).
notation < "\infrule hbox(\emsp ab \emsp) mstyle color #ff0000 (a) mstyle color #ff0000 (∧\sub(\e_\r))" with precedence 19
for @{ 'And_elim_r_ko_2 $ab $a }.
interpretation "And_elim_r_ko_2" 'And_elim_r_ko_2 ab a =
- (cast _ _ (show a (cast _ _ (And_elim_r _ _ (cast _ _ ab))))).
+ (cast ? ? (show a (cast ? ? (And_elim_r ? ? (cast ? ? ab))))).
notation < "maction (\infrule hbox(\emsp ab \emsp) a (∧\sub(\e_\r))) (\vdots)" with precedence 19
for @{ 'And_elim_r_ok_1 $ab $a }.
interpretation "And_elim_r_ok_1" 'And_elim_r_ok_1 ab a =
- (show a (And_elim_r _ _ ab)).
+ (show a (And_elim_r ? ? ab)).
notation < "\infrule hbox(\emsp ab \emsp) mstyle color #ff0000 (a) (∧\sub(\e_\r))" with precedence 19
for @{ 'And_elim_r_ok_2 $ab $a }.
interpretation "And_elim_r_ok_2" 'And_elim_r_ok_2 ab a =
- (cast _ _ (show a (And_elim_r _ _ ab))).
+ (cast ? ? (show a (And_elim_r ? ? ab))).
-notation > "∧_'e_r' term 90 ab" with precedence 19
+notation > "∧#'e_r' term 90 ab" with precedence 19
for @{ 'And_elim_r (show $ab ?) }.
-interpretation "And_elim_r KO" 'And_elim_r a = (cast _ _ (And_elim_r _ _ (cast (And unit unit) _ a))).
-interpretation "And_elim_r OK" 'And_elim_r a = (And_elim_r _ _ a).
+interpretation "And_elim_r KO" 'And_elim_r a = (cast ? ? (And_elim_r ? ? (cast (And unit unit) ? a))).
+interpretation "And_elim_r OK" 'And_elim_r a = (And_elim_r ? ? a).
(* ∨ introduction *)
notation < "\infrule hbox(\emsp a \emsp) ab mstyle color #ff0000 (∨\sub(\i_\l))" with precedence 19
for @{ 'Or_intro_l_ko_1 $a $ab }.
interpretation "Or_intro_l_ko_1" 'Or_intro_l_ko_1 a ab =
- (show ab (cast _ _ (Or_intro_l _ _ a))).
+ (show ab (cast ? ? (Or_intro_l ? ? a))).
notation < "\infrule hbox(\emsp a \emsp) mstyle color #ff0000 (ab) mstyle color #ff0000 (∨\sub(\i_\l))" with precedence 19
for @{ 'Or_intro_l_ko_2 $a $ab }.
interpretation "Or_intro_l_ko_2" 'Or_intro_l_ko_2 a ab =
- (cast _ _ (show ab (cast _ _ (Or_intro_l _ _ a)))).
+ (cast ? ? (show ab (cast ? ? (Or_intro_l ? ? a)))).
notation < "maction (\infrule hbox(\emsp a \emsp) ab (∨\sub(\i_\l))) (\vdots)" with precedence 19
for @{ 'Or_intro_l_ok_1 $a $ab }.
interpretation "Or_intro_l_ok_1" 'Or_intro_l_ok_1 a ab =
- (show ab (Or_intro_l _ _ a)).
+ (show ab (Or_intro_l ? ? a)).
notation < "\infrule hbox(\emsp a \emsp) mstyle color #ff0000 (ab) (∨\sub(\i_\l))" with precedence 19
for @{ 'Or_intro_l_ok_2 $a $ab }.
interpretation "Or_intro_l_ok_2" 'Or_intro_l_ok_2 a ab =
- (cast _ _ (show ab (Or_intro_l _ _ a))).
+ (cast ? ? (show ab (Or_intro_l ? ? a))).
-notation > "∨_'i_l' term 90 a" with precedence 19
+notation > "∨#'i_l' term 90 a" with precedence 19
for @{ 'Or_intro_l (show $a ?) }.
-interpretation "Or_intro_l KO" 'Or_intro_l a = (cast _ (Or _ unit) (Or_intro_l _ _ a)).
-interpretation "Or_intro_l OK" 'Or_intro_l a = (Or_intro_l _ _ a).
+interpretation "Or_intro_l KO" 'Or_intro_l a = (cast ? (Or ? unit) (Or_intro_l ? ? a)).
+interpretation "Or_intro_l OK" 'Or_intro_l a = (Or_intro_l ? ? a).
notation < "\infrule hbox(\emsp a \emsp) ab mstyle color #ff0000 (∨\sub(\i_\r))" with precedence 19
for @{ 'Or_intro_r_ko_1 $a $ab }.
interpretation "Or_intro_r_ko_1" 'Or_intro_r_ko_1 a ab =
- (show ab (cast _ _ (Or_intro_r _ _ a))).
+ (show ab (cast ? ? (Or_intro_r ? ? a))).
notation < "\infrule hbox(\emsp a \emsp) mstyle color #ff0000 (ab) mstyle color #ff0000 (∨\sub(\i_\r))" with precedence 19
for @{ 'Or_intro_r_ko_2 $a $ab }.
interpretation "Or_intro_r_ko_2" 'Or_intro_r_ko_2 a ab =
- (cast _ _ (show ab (cast _ _ (Or_intro_r _ _ a)))).
+ (cast ? ? (show ab (cast ? ? (Or_intro_r ? ? a)))).
notation < "maction (\infrule hbox(\emsp a \emsp) ab (∨\sub(\i_\r))) (\vdots)" with precedence 19
for @{ 'Or_intro_r_ok_1 $a $ab }.
interpretation "Or_intro_r_ok_1" 'Or_intro_r_ok_1 a ab =
- (show ab (Or_intro_r _ _ a)).
+ (show ab (Or_intro_r ? ? a)).
notation < "\infrule hbox(\emsp a \emsp) mstyle color #ff0000 (ab) (∨\sub(\i_\r))" with precedence 19
for @{ 'Or_intro_r_ok_2 $a $ab }.
interpretation "Or_intro_r_ok_2" 'Or_intro_r_ok_2 a ab =
- (cast _ _ (show ab (Or_intro_r _ _ a))).
+ (cast ? ? (show ab (Or_intro_r ? ? a))).
-notation > "∨_'i_r' term 90 a" with precedence 19
+notation > "∨#'i_r' term 90 a" with precedence 19
for @{ 'Or_intro_r (show $a ?) }.
-interpretation "Or_intro_r KO" 'Or_intro_r a = (cast _ (Or unit _) (Or_intro_r _ _ a)).
-interpretation "Or_intro_r OK" 'Or_intro_r a = (Or_intro_r _ _ a).
+interpretation "Or_intro_r KO" 'Or_intro_r a = (cast ? (Or unit ?) (Or_intro_r ? ? a)).
+interpretation "Or_intro_r OK" 'Or_intro_r a = (Or_intro_r ? ? a).
(* ∨ elimination *)
notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp ac \emsp\emsp\emsp bc \emsp) c (mstyle color #ff0000 (∨\sub\e \emsp) ident Ha \emsp ident Hb)" with precedence 19
for @{ 'Or_elim_ko_1 $ab $c (λ${ident Ha}:$ta.$ac) (λ${ident Hb}:$tb.$bc) }.
interpretation "Or_elim_ko_1" 'Or_elim_ko_1 ab c \eta.ac \eta.bc =
- (show c (cast _ _ (Or_elim _ _ _ (cast _ _ ab) (cast _ _ ac) (cast _ _ bc)))).
+ (show c (cast ? ? (Or_elim ? ? ? (cast ? ? ab) (cast ? ? ac) (cast ? ? bc)))).
notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp ac \emsp\emsp\emsp bc \emsp) mstyle color #ff0000 (c) (mstyle color #ff0000 (∨\sub\e) \emsp ident Ha \emsp ident Hb)" with precedence 19
for @{ 'Or_elim_ko_2 $ab (λ${ident Ha}:$ta.$ac) (λ${ident Hb}:$tb.$bc) $c }.
interpretation "Or_elim_ko_2" 'Or_elim_ko_2 ab \eta.ac \eta.bc c =
- (cast _ _ (show c (cast _ _ (Or_elim _ _ _ (cast _ _ ab) (cast _ _ ac) (cast _ _ bc))))).
+ (cast ? ? (show c (cast ? ? (Or_elim ? ? ? (cast ? ? ab) (cast ? ? ac) (cast ? ? bc))))).
notation < "maction (\infrule hbox(\emsp ab \emsp\emsp\emsp ac \emsp\emsp\emsp bc \emsp) c (∨\sub\e \emsp ident Ha \emsp ident Hb)) (\vdots)" with precedence 19
for @{ 'Or_elim_ok_1 $ab (λ${ident Ha}:$ta.$ac) (λ${ident Hb}:$tb.$bc) $c }.
interpretation "Or_elim_ok_1" 'Or_elim_ok_1 ab \eta.ac \eta.bc c =
- (show c (Or_elim _ _ _ ab ac bc)).
+ (show c (Or_elim ? ? ? ab ac bc)).
notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp ac \emsp\emsp\emsp bc \emsp) mstyle color #ff0000 (c) (∨\sub\e \emsp ident Ha \emsp ident Hb)" with precedence 19
for @{ 'Or_elim_ok_2 $ab (λ${ident Ha}:$ta.$ac) (λ${ident Hb}:$tb.$bc) $c }.
interpretation "Or_elim_ok_2" 'Or_elim_ok_2 ab \eta.ac \eta.bc c =
- (cast _ _ (show c (Or_elim _ _ _ ab ac bc))).
+ (cast ? ? (show c (Or_elim ? ? ? ab ac bc))).
definition unit_to ≝ λx:CProp.unit → x.
-notation > "∨_'e' term 90 ab [ident Ha] term 90 cl [ident Hb] term 90 cr" with precedence 19
+notation > "∨#'e' term 90 ab [ident Ha] term 90 cl [ident Hb] term 90 cr" with precedence 19
for @{ 'Or_elim (show $ab ?) (λ${ident Ha}.show $cl ?) (λ${ident Hb}.show $cr ?) }.
interpretation "Or_elim KO" 'Or_elim ab ac bc =
- (cast _ _ (Or_elim _ _ _
- (cast (Or unit unit) _ ab)
- (cast (unit_to unit) (unit_to _) ac)
- (cast (unit_to unit) (unit_to _) bc))).
-interpretation "Or_elim OK" 'Or_elim ab ac bc = (Or_elim _ _ _ ab ac bc).
+ (cast ? ? (Or_elim ? ? ?
+ (cast (Or unit unit) ? ab)
+ (cast (unit_to unit) (unit_to ?) ac)
+ (cast (unit_to unit) (unit_to ?) bc))).
+interpretation "Or_elim OK" 'Or_elim ab ac bc = (Or_elim ? ? ? ab ac bc).
(* ⊤ introduction *)
notation < "\infrule \nbsp ⊤ mstyle color #ff0000 (⊤\sub\i)" with precedence 19
for @{'Top_intro_ko_1}.
interpretation "Top_intro_ko_1" 'Top_intro_ko_1 =
- (show _ (cast _ _ Top_intro)).
+ (show ? (cast ? ? Top_intro)).
notation < "\infrule \nbsp mstyle color #ff0000 (⊤) mstyle color #ff0000 (⊤\sub\i)" with precedence 19
for @{'Top_intro_ko_2}.
interpretation "Top_intro_ko_2" 'Top_intro_ko_2 =
- (cast _ _ (show _ (cast _ _ Top_intro))).
+ (cast ? ? (show ? (cast ? ? Top_intro))).
notation < "maction (\infrule \nbsp ⊤ (⊤\sub\i)) (\vdots)" with precedence 19
for @{'Top_intro_ok_1}.
-interpretation "Top_intro_ok_1" 'Top_intro_ok_1 = (show _ Top_intro).
+interpretation "Top_intro_ok_1" 'Top_intro_ok_1 = (show ? Top_intro).
notation < "maction (\infrule \nbsp ⊤ (⊤\sub\i)) (\vdots)" with precedence 19
for @{'Top_intro_ok_2 }.
-interpretation "Top_intro_ok_2" 'Top_intro_ok_2 = (cast _ _ (show _ Top_intro)).
+interpretation "Top_intro_ok_2" 'Top_intro_ok_2 = (cast ? ? (show ? Top_intro)).
-notation > "⊤_'i'" with precedence 19 for @{ 'Top_intro }.
-interpretation "Top_intro KO" 'Top_intro = (cast _ _ Top_intro).
+notation > "⊤#'i'" with precedence 19 for @{ 'Top_intro }.
+interpretation "Top_intro KO" 'Top_intro = (cast ? ? Top_intro).
interpretation "Top_intro OK" 'Top_intro = Top_intro.
(* ⊥ introduction *)
notation < "\infrule b a mstyle color #ff0000 (⊥\sub\e)" with precedence 19
for @{'Bot_elim_ko_1 $a $b}.
interpretation "Bot_elim_ko_1" 'Bot_elim_ko_1 a b =
- (show a (Bot_elim _ (cast _ _ b))).
+ (show a (Bot_elim ? (cast ? ? b))).
notation < "\infrule b mstyle color #ff0000 (a) mstyle color #ff0000 (⊥\sub\e)" with precedence 19
for @{'Bot_elim_ko_2 $a $b}.
interpretation "Bot_elim_ko_2" 'Bot_elim_ko_2 a b =
- (cast _ _ (show a (Bot_elim _ (cast _ _ b)))).
+ (cast ? ? (show a (Bot_elim ? (cast ? ? b)))).
notation < "maction (\infrule b a (⊥\sub\e)) (\vdots)" with precedence 19
for @{'Bot_elim_ok_1 $a $b}.
interpretation "Bot_elim_ok_1" 'Bot_elim_ok_1 a b =
- (show a (Bot_elim _ b)).
+ (show a (Bot_elim ? b)).
notation < "\infrule b mstyle color #ff0000 (a) (⊥\sub\e)" with precedence 19
for @{'Bot_elim_ok_2 $a $b}.
interpretation "Bot_elim_ok_2" 'Bot_elim_ok_2 a b =
- (cast _ _ (show a (Bot_elim _ b))).
+ (cast ? ? (show a (Bot_elim ? b))).
-notation > "⊥_'e' term 90 b" with precedence 19
+notation > "⊥#'e' term 90 b" with precedence 19
for @{ 'Bot_elim (show $b ?) }.
-interpretation "Bot_elim KO" 'Bot_elim a = (Bot_elim _ (cast _ _ a)).
-interpretation "Bot_elim OK" 'Bot_elim a = (Bot_elim _ a).
+interpretation "Bot_elim KO" 'Bot_elim a = (Bot_elim ? (cast ? ? a)).
+interpretation "Bot_elim OK" 'Bot_elim a = (Bot_elim ? a).
(* ¬ introduction *)
notation < "\infrule hbox(\emsp b \emsp) ab (mstyle color #ff0000 (\lnot\sub(\emsp\i)) \emsp ident H)" with precedence 19
for @{ 'Not_intro_ko_1 $ab (λ${ident H}:$p.$b) }.
interpretation "Not_intro_ko_1" 'Not_intro_ko_1 ab \eta.b =
- (show ab (cast _ _ (Not_intro _ (cast _ _ b)))).
+ (show ab (cast ? ? (Not_intro ? (cast ? ? b)))).
notation < "\infrule hbox(\emsp b \emsp) mstyle color #ff0000 (ab) (mstyle color #ff0000 (\lnot\sub(\emsp\i)) \emsp ident H)" with precedence 19
for @{ 'Not_intro_ko_2 $ab (λ${ident H}:$p.$b) }.
interpretation "Not_intro_ko_2" 'Not_intro_ko_2 ab \eta.b =
- (cast _ _ (show ab (cast _ _ (Not_intro _ (cast _ _ b))))).
+ (cast ? ? (show ab (cast ? ? (Not_intro ? (cast ? ? b))))).
notation < "maction (\infrule hbox(\emsp b \emsp) ab (\lnot\sub(\emsp\i) \emsp ident H) ) (\vdots)" with precedence 19
for @{ 'Not_intro_ok_1 $ab (λ${ident H}:$p.$b) }.
interpretation "Not_intro_ok_1" 'Not_intro_ok_1 ab \eta.b =
- (show ab (Not_intro _ b)).
+ (show ab (Not_intro ? b)).
notation < "\infrule hbox(\emsp b \emsp) mstyle color #ff0000 (ab) (\lnot\sub(\emsp\i) \emsp ident H) " with precedence 19
for @{ 'Not_intro_ok_2 $ab (λ${ident H}:$p.$b) }.
interpretation "Not_intro_ok_2" 'Not_intro_ok_2 ab \eta.b =
- (cast _ _ (show ab (Not_intro _ b))).
+ (cast ? ? (show ab (Not_intro ? b))).
-notation > "¬_'i' [ident H] term 90 b" with precedence 19
+notation > "¬#'i' [ident H] term 90 b" with precedence 19
for @{ 'Not_intro (λ${ident H}.show $b ?) }.
-interpretation "Not_intro KO" 'Not_intro a = (cast _ _ (Not_intro _ (cast _ _ a))).
-interpretation "Not_intro OK" 'Not_intro a = (Not_intro _ a).
+interpretation "Not_intro KO" 'Not_intro a = (cast ? ? (Not_intro ? (cast ? ? a))).
+interpretation "Not_intro OK" 'Not_intro a = (Not_intro ? a).
(* ¬ elimination *)
notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) b mstyle color #ff0000 (\lnot\sub(\emsp\e)) " with precedence 19
for @{ 'Not_elim_ko_1 $ab $a $b }.
interpretation "Not_elim_ko_1" 'Not_elim_ko_1 ab a b =
- (show b (cast _ _ (Not_elim _ (cast _ _ ab) (cast _ _ a)))).
+ (show b (cast ? ? (Not_elim ? (cast ? ? ab) (cast ? ? a)))).
notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) mstyle color #ff0000 (b) mstyle color #ff0000 (\lnot\sub(\emsp\e)) " with precedence 19
for @{ 'Not_elim_ko_2 $ab $a $b }.
interpretation "Not_elim_ko_2" 'Not_elim_ko_2 ab a b =
- (cast _ _ (show b (cast _ _ (Not_elim _ (cast _ _ ab) (cast _ _ a))))).
+ (cast ? ? (show b (cast ? ? (Not_elim ? (cast ? ? ab) (cast ? ? a))))).
notation < "maction (\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) b (\lnot\sub(\emsp\e)) ) (\vdots)" with precedence 19
for @{ 'Not_elim_ok_1 $ab $a $b }.
interpretation "Not_elim_ok_1" 'Not_elim_ok_1 ab a b =
- (show b (Not_elim _ ab a)).
+ (show b (Not_elim ? ab a)).
notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) mstyle color #ff0000 (b) (\lnot\sub(\emsp\e)) " with precedence 19
for @{ 'Not_elim_ok_2 $ab $a $b }.
interpretation "Not_elim_ok_2" 'Not_elim_ok_2 ab a b =
- (cast _ _ (show b (Not_elim _ ab a))).
+ (cast ? ? (show b (Not_elim ? ab a))).
-notation > "¬_'e' term 90 ab term 90 a" with precedence 19
+notation > "¬#'e' term 90 ab term 90 a" with precedence 19
for @{ 'Not_elim (show $ab ?) (show $a ?) }.
interpretation "Not_elim KO" 'Not_elim ab a =
- (cast _ _ (Not_elim unit (cast _ _ ab) (cast _ _ a))).
+ (cast ? ? (Not_elim unit (cast ? ? ab) (cast ? ? a))).
interpretation "Not_elim OK" 'Not_elim ab a =
- (Not_elim _ ab a).
+ (Not_elim ? ab a).
(* RAA *)
notation < "\infrule hbox(\emsp Px \emsp) Pn (mstyle color #ff0000 (\RAA) \emsp ident x)" with precedence 19
for @{ 'RAA_ko_1 (λ${ident x}:$tx.$Px) $Pn }.
interpretation "RAA_ko_1" 'RAA_ko_1 Px Pn =
- (show Pn (cast _ _ (Raa _ (cast _ _ Px)))).
+ (show Pn (cast ? ? (Raa ? (cast ? ? Px)))).
notation < "\infrule hbox(\emsp Px \emsp) mstyle color #ff0000 (Pn) (mstyle color #ff0000 (\RAA) \emsp ident x)" with precedence 19
for @{ 'RAA_ko_2 (λ${ident x}:$tx.$Px) $Pn }.
interpretation "RAA_ko_2" 'RAA_ko_2 Px Pn =
- (cast _ _ (show Pn (cast _ _ (Raa _ (cast _ _ Px))))).
+ (cast ? ? (show Pn (cast ? ? (Raa ? (cast ? ? Px))))).
notation < "maction (\infrule hbox(\emsp Px \emsp) Pn (\RAA \emsp ident x)) (\vdots)" with precedence 19
for @{ 'RAA_ok_1 (λ${ident x}:$tx.$Px) $Pn }.
interpretation "RAA_ok_1" 'RAA_ok_1 Px Pn =
- (show Pn (Raa _ Px)).
+ (show Pn (Raa ? Px)).
notation < "\infrule hbox(\emsp Px \emsp) mstyle color #ff0000 (Pn) (\RAA \emsp ident x)" with precedence 19
for @{ 'RAA_ok_2 (λ${ident x}:$tx.$Px) $Pn }.
interpretation "RAA_ok_2" 'RAA_ok_2 Px Pn =
- (cast _ _ (show Pn (Raa _ Px))).
+ (cast ? ? (show Pn (Raa ? Px))).
notation > "'RAA' [ident H] term 90 b" with precedence 19
for @{ 'Raa (λ${ident H}.show $b ?) }.
-interpretation "RAA KO" 'Raa p = (cast _ unit (Raa _ (cast _ (unit_to _) p))).
-interpretation "RAA OK" 'Raa p = (Raa _ p).
+interpretation "RAA KO" 'Raa p = (cast ? unit (Raa ? (cast ? (unit_to ?) p))).
+interpretation "RAA OK" 'Raa p = (Raa ? p).
(* ∃ introduction *)
notation < "\infrule hbox(\emsp Pn \emsp) Px mstyle color #ff0000 (∃\sub\i)" with precedence 19
for @{ 'Exists_intro_ko_1 $Pn $Px }.
interpretation "Exists_intro_ko_1" 'Exists_intro_ko_1 Pn Px =
- (show Px (cast _ _ (Exists_intro _ _ _ (cast _ _ Pn)))).
+ (show Px (cast ? ? (Exists_intro ? ? ? (cast ? ? Pn)))).
notation < "\infrule hbox(\emsp Pn \emsp) mstyle color #ff0000 (Px) mstyle color #ff0000 (∃\sub\i)" with precedence 19
for @{ 'Exists_intro_ko_2 $Pn $Px }.
interpretation "Exists_intro_ko_2" 'Exists_intro_ko_2 Pn Px =
- (cast _ _ (show Px (cast _ _ (Exists_intro _ _ _ (cast _ _ Pn))))).
+ (cast ? ? (show Px (cast ? ? (Exists_intro ? ? ? (cast ? ? Pn))))).
notation < "maction (\infrule hbox(\emsp Pn \emsp) Px (∃\sub\i)) (\vdots)" with precedence 19
for @{ 'Exists_intro_ok_1 $Pn $Px }.
interpretation "Exists_intro_ok_1" 'Exists_intro_ok_1 Pn Px =
- (show Px (Exists_intro _ _ _ Pn)).
+ (show Px (Exists_intro ? ? ? Pn)).
notation < "\infrule hbox(\emsp Pn \emsp) mstyle color #ff0000 (Px) (∃\sub\i)" with precedence 19
for @{ 'Exists_intro_ok_2 $Pn $Px }.
interpretation "Exists_intro_ok_2" 'Exists_intro_ok_2 Pn Px =
- (cast _ _ (show Px (Exists_intro _ _ _ Pn))).
+ (cast ? ? (show Px (Exists_intro ? ? ? Pn))).
-notation >"∃_'i' {term 90 t} term 90 Pt" non associative with precedence 19
+notation >"∃#'i' {term 90 t} term 90 Pt" non associative with precedence 19
for @{'Exists_intro $t (λw.? w) (show $Pt ?)}.
interpretation "Exists_intro KO" 'Exists_intro t P Pt =
- (cast _ _ (Exists_intro sort P t (cast _ _ Pt))).
+ (cast ? ? (Exists_intro sort P t (cast ? ? Pt))).
interpretation "Exists_intro OK" 'Exists_intro t P Pt =
(Exists_intro sort P t Pt).
notation < "\infrule hbox(\emsp ExPx \emsp\emsp\emsp Pc \emsp) c (mstyle color #ff0000 (∃\sub\e) \emsp ident n \emsp ident HPn)" with precedence 19
for @{ 'Exists_elim_ko_1 $ExPx (λ${ident n}:$tn.λ${ident HPn}:$Pn.$Pc) $c }.
interpretation "Exists_elim_ko_1" 'Exists_elim_ko_1 ExPx Pc c =
- (show c (cast _ _ (Exists_elim _ _ _ (cast _ _ ExPx) (cast _ _ Pc)))).
+ (show c (cast ? ? (Exists_elim ? ? ? (cast ? ? ExPx) (cast ? ? Pc)))).
notation < "\infrule hbox(\emsp ExPx \emsp\emsp\emsp Pc \emsp) mstyle color #ff0000 (c) (mstyle color #ff0000 (∃\sub\e) \emsp ident n \emsp ident HPn)" with precedence 19
for @{ 'Exists_elim_ko_2 $ExPx (λ${ident n}:$tn.λ${ident HPn}:$Pn.$Pc) $c }.
interpretation "Exists_elim_ko_2" 'Exists_elim_ko_2 ExPx Pc c =
- (cast _ _ (show c (cast _ _ (Exists_elim _ _ _ (cast _ _ ExPx) (cast _ _ Pc))))).
+ (cast ? ? (show c (cast ? ? (Exists_elim ? ? ? (cast ? ? ExPx) (cast ? ? Pc))))).
notation < "maction (\infrule hbox(\emsp ExPx \emsp\emsp\emsp Pc \emsp) c (∃\sub\e \emsp ident n \emsp ident HPn)) (\vdots)" with precedence 19
for @{ 'Exists_elim_ok_1 $ExPx (λ${ident n}:$tn.λ${ident HPn}:$Pn.$Pc) $c }.
interpretation "Exists_elim_ok_1" 'Exists_elim_ok_1 ExPx Pc c =
- (show c (Exists_elim _ _ _ ExPx Pc)).
+ (show c (Exists_elim ? ? ? ExPx Pc)).
notation < "\infrule hbox(\emsp ExPx \emsp\emsp\emsp Pc \emsp) mstyle color #ff0000 (c) (∃\sub\e \emsp ident n \emsp ident HPn)" with precedence 19
for @{ 'Exists_elim_ok_2 $ExPx (λ${ident n}:$tn.λ${ident HPn}:$Pn.$Pc) $c }.
interpretation "Exists_elim_ok_2" 'Exists_elim_ok_2 ExPx Pc c =
- (cast _ _ (show c (Exists_elim _ _ _ ExPx Pc))).
+ (cast ? ? (show c (Exists_elim ? ? ? ExPx Pc))).
definition ex_concl := λx:sort → CProp.∀y:sort.unit → x y.
definition ex_concl_dummy := ∀y:sort.unit → unit.
definition fake_pred := λx:sort.unit.
-notation >"∃_'e' term 90 ExPt {ident t} [ident H] term 90 c" non associative with precedence 19
+notation >"∃#'e' term 90 ExPt {ident t} [ident H] term 90 c" non associative with precedence 19
for @{'Exists_elim (λx.? x) (show $ExPt ?) (λ${ident t}:sort.λ${ident H}.show $c ?)}.
interpretation "Exists_elim KO" 'Exists_elim P ExPt c =
- (cast _ _ (Exists_elim sort P _
- (cast (Exists _ P) _ ExPt)
- (cast ex_concl_dummy (ex_concl _) c))).
+ (cast ? ? (Exists_elim sort P ?
+ (cast (Exists ? P) ? ExPt)
+ (cast ex_concl_dummy (ex_concl ?) c))).
interpretation "Exists_elim OK" 'Exists_elim P ExPt c =
- (Exists_elim sort P _ ExPt c).
+ (Exists_elim sort P ? ExPt c).
(* ∀ introduction *)
notation < "\infrule hbox(\emsp Px \emsp) Pn (mstyle color #ff0000 (∀\sub\i) \emsp ident x)" with precedence 19
for @{ 'Forall_intro_ko_1 (λ${ident x}:$tx.$Px) $Pn }.
interpretation "Forall_intro_ko_1" 'Forall_intro_ko_1 Px Pn =
- (show Pn (cast _ _ (Forall_intro _ _ (cast _ _ Px)))).
+ (show Pn (cast ? ? (Forall_intro ? ? (cast ? ? Px)))).
notation < "\infrule hbox(\emsp Px \emsp) mstyle color #ff0000(Pn) (mstyle color #ff0000 (∀\sub\i) \emsp ident x)" with precedence 19
for @{ 'Forall_intro_ko_2 (λ${ident x}:$tx.$Px) $Pn }.
interpretation "Forall_intro_ko_2" 'Forall_intro_ko_2 Px Pn =
- (cast _ _ (show Pn (cast _ _ (Forall_intro _ _ (cast _ _ Px))))).
+ (cast ? ? (show Pn (cast ? ? (Forall_intro ? ? (cast ? ? Px))))).
notation < "maction (\infrule hbox(\emsp Px \emsp) Pn (∀\sub\i \emsp ident x)) (\vdots)" with precedence 19
for @{ 'Forall_intro_ok_1 (λ${ident x}:$tx.$Px) $Pn }.
interpretation "Forall_intro_ok_1" 'Forall_intro_ok_1 Px Pn =
- (show Pn (Forall_intro _ _ Px)).
+ (show Pn (Forall_intro ? ? Px)).
notation < "\infrule hbox(\emsp Px \emsp) mstyle color #ff0000 (Pn) (∀\sub\i \emsp ident x)" with precedence 19
for @{ 'Forall_intro_ok_2 (λ${ident x}:$tx.$Px) $Pn }.
interpretation "Forall_intro_ok_2" 'Forall_intro_ok_2 Px Pn =
- (cast _ _ (show Pn (Forall_intro _ _ Px))).
+ (cast ? ? (show Pn (Forall_intro ? ? Px))).
-notation > "∀_'i' {ident y} term 90 Px" non associative with precedence 19
+notation > "∀#'i' {ident y} term 90 Px" non associative with precedence 19
for @{ 'Forall_intro (λ_.?) (λ${ident y}.show $Px ?) }.
interpretation "Forall_intro KO" 'Forall_intro P Px =
- (cast _ _ (Forall_intro sort P (cast _ _ Px))).
+ (cast ? ? (Forall_intro sort P (cast ? ? Px))).
interpretation "Forall_intro OK" 'Forall_intro P Px =
(Forall_intro sort P Px).
notation < "\infrule hbox(\emsp Px \emsp) Pn (mstyle color #ff0000 (∀\sub\e))" with precedence 19
for @{ 'Forall_elim_ko_1 $Px $Pn }.
interpretation "Forall_elim_ko_1" 'Forall_elim_ko_1 Px Pn =
- (show Pn (cast _ _ (Forall_elim _ _ _ (cast _ _ Px)))).
+ (show Pn (cast ? ? (Forall_elim ? ? ? (cast ? ? Px)))).
notation < "\infrule hbox(\emsp Px \emsp) mstyle color #ff0000(Pn) (mstyle color #ff0000 (∀\sub\e))" with precedence 19
for @{ 'Forall_elim_ko_2 $Px $Pn }.
interpretation "Forall_elim_ko_2" 'Forall_elim_ko_2 Px Pn =
- (cast _ _ (show Pn (cast _ _ (Forall_elim _ _ _ (cast _ _ Px))))).
+ (cast ? ? (show Pn (cast ? ? (Forall_elim ? ? ? (cast ? ? Px))))).
notation < "maction (\infrule hbox(\emsp Px \emsp) Pn (∀\sub\e)) (\vdots)" with precedence 19
for @{ 'Forall_elim_ok_1 $Px $Pn }.
interpretation "Forall_elim_ok_1" 'Forall_elim_ok_1 Px Pn =
- (show Pn (Forall_elim _ _ _ Px)).
+ (show Pn (Forall_elim ? ? ? Px)).
notation < "\infrule hbox(\emsp Px \emsp) mstyle color #ff0000 (Pn) (∀\sub\e)" with precedence 19
for @{ 'Forall_elim_ok_2 $Px $Pn }.
interpretation "Forall_elim_ok_2" 'Forall_elim_ok_2 Px Pn =
- (cast _ _ (show Pn (Forall_elim _ _ _ Px))).
+ (cast ? ? (show Pn (Forall_elim ? ? ? Px))).
-notation > "∀_'e' {term 90 t} term 90 Pn" non associative with precedence 19
+notation > "∀#'e' {term 90 t} term 90 Pn" non associative with precedence 19
for @{ 'Forall_elim (λ_.?) $t (show $Pn ?) }.
interpretation "Forall_elim KO" 'Forall_elim P t Px =
- (cast _ unit (Forall_elim sort P t (cast _ _ Px))).
+ (cast ? unit (Forall_elim sort P t (cast ? ? Px))).
interpretation "Forall_elim OK" 'Forall_elim P t Px =
(Forall_elim sort P t Px).
(* already proved lemma *)
-definition hide_args : ∀A:Type.∀a:A.A := λA:Type.λa:A.a.
+definition hide_args : ∀A:Type.A→A := λA:Type.λa:A.a.
notation < "t" non associative with precedence 90 for @{'hide_args $t}.
-interpretation "hide 0 args" 'hide_args t = (hide_args _ t).
-interpretation "hide 1 args" 'hide_args t = (hide_args _ t _).
-interpretation "hide 2 args" 'hide_args t = (hide_args _ t _ _).
-interpretation "hide 3 args" 'hide_args t = (hide_args _ t _ _ _).
-interpretation "hide 4 args" 'hide_args t = (hide_args _ t _ _ _ _).
-interpretation "hide 5 args" 'hide_args t = (hide_args _ t _ _ _ _ _).
-interpretation "hide 6 args" 'hide_args t = (hide_args _ t _ _ _ _ _ _).
-interpretation "hide 7 args" 'hide_args t = (hide_args _ t _ _ _ _ _ _ _).
+interpretation "hide 0 args" 'hide_args t = (hide_args ? t).
+interpretation "hide 1 args" 'hide_args t = (hide_args ? t ?).
+interpretation "hide 2 args" 'hide_args t = (hide_args ? t ? ?).
+interpretation "hide 3 args" 'hide_args t = (hide_args ? t ? ? ?).
+interpretation "hide 4 args" 'hide_args t = (hide_args ? t ? ? ? ?).
+interpretation "hide 5 args" 'hide_args t = (hide_args ? t ? ? ? ? ?).
+interpretation "hide 6 args" 'hide_args t = (hide_args ? t ? ? ? ? ? ?).
+interpretation "hide 7 args" 'hide_args t = (hide_args ? t ? ? ? ? ? ? ?).
(* more args crashes the pattern matcher *)
non associative with precedence 19
for @{ 'lemma_ko_1 $p ($H : $_) }.
interpretation "lemma_ko_1" 'lemma_ko_1 p H =
- (show p (cast _ _ (Lemma _ (cast _ _ H)))).
+ (show p (cast ? ? (Lemma ? (cast ? ? H)))).
notation < "\infrule
(\infrule
non associative with precedence 19
for @{ 'lemma_ko_2 $p ($H : $_) }.
interpretation "lemma_ko_2" 'lemma_ko_2 p H =
- (cast _ _ (show p (cast _ _
- (Lemma _ (cast _ _ H))))).
+ (cast ? ? (show p (cast ? ?
+ (Lemma ? (cast ? ? H))))).
notation < "\infrule
non associative with precedence 19
for @{ 'lemma_ok_1 $p ($H : $_) }.
interpretation "lemma_ok_1" 'lemma_ok_1 p H =
- (show p (Lemma _ H)).
+ (show p (Lemma ? H)).
notation < "\infrule
(\infrule
non associative with precedence 19
for @{ 'lemma_ok_2 $p ($H : $_) }.
interpretation "lemma_ok_2" 'lemma_ok_2 p H =
- (cast _ _ (show p (Lemma _ H))).
+ (cast ? ? (show p (Lemma ? H))).
notation > "'lem' 0 term 90 l" non associative with precedence 19
for @{ 'Lemma (hide_args ? $l : ?) }.
interpretation "lemma KO" 'Lemma l =
- (cast _ _ (Lemma unit (cast unit _ l))).
-interpretation "lemma OK" 'Lemma l = (Lemma _ l).
+ (cast ? ? (Lemma unit (cast unit ? l))).
+interpretation "lemma OK" 'Lemma l = (Lemma ? l).
(* already proved lemma, 1 assumption *)
non associative with precedence 19
for @{ 'lemma1_ko_1 $a $p ($H : $_) }.
interpretation "lemma1_ko_1" 'lemma1_ko_1 a p H =
- (show p (cast _ _ (Lemma1 _ _ (cast _ _ H) (cast _ _ a)))).
+ (show p (cast ? ? (Lemma1 ? ? (cast ? ? H) (cast ? ? a)))).
notation < "\infrule
(\infrule
non associative with precedence 19
for @{ 'lemma1_ko_2 $a $p ($H : $_) }.
interpretation "lemma1_ko_2" 'lemma1_ko_2 a p H =
- (cast _ _ (show p (cast _ _
- (Lemma1 _ _ (cast _ _ H) (cast _ _ a))))).
+ (cast ? ? (show p (cast ? ?
+ (Lemma1 ? ? (cast ? ? H) (cast ? ? a))))).
notation < "\infrule
non associative with precedence 19
for @{ 'lemma1_ok_1 $a $p ($H : $_) }.
interpretation "lemma1_ok_1" 'lemma1_ok_1 a p H =
- (show p (Lemma1 _ _ H a)).
+ (show p (Lemma1 ? ? H a)).
notation < "\infrule
(\infrule
non associative with precedence 19
for @{ 'lemma1_ok_2 $a $p ($H : $_) }.
interpretation "lemma1_ok_2" 'lemma1_ok_2 a p H =
- (cast _ _ (show p (Lemma1 _ _ H a))).
+ (cast ? ? (show p (Lemma1 ? ? H a))).
notation > "'lem' 1 term 90 l term 90 p" non associative with precedence 19
for @{ 'Lemma1 (hide_args ? $l : ?) (show $p ?) }.
interpretation "lemma 1 KO" 'Lemma1 l p =
- (cast _ _ (Lemma1 unit unit (cast (Imply unit unit) _ l) (cast unit _ p))).
-interpretation "lemma 1 OK" 'Lemma1 l p = (Lemma1 _ _ l p).
+ (cast ? ? (Lemma1 unit unit (cast (Imply unit unit) ? l) (cast unit ? p))).
+interpretation "lemma 1 OK" 'Lemma1 l p = (Lemma1 ? ? l p).
(* already proved lemma, 2 assumptions *)
definition Lemma2 : ΠA,B,C. (A ⇒ B ⇒ C) → A → B → C ≝
non associative with precedence 19
for @{ 'lemma2_ko_1 $a $b $p ($H : $_) }.
interpretation "lemma2_ko_1" 'lemma2_ko_1 a b p H =
- (show p (cast _ _ (Lemma2 _ _ _ (cast _ _ H) (cast _ _ a) (cast _ _ b)))).
+ (show p (cast ? ? (Lemma2 ? ? ? (cast ? ? H) (cast ? ? a) (cast ? ? b)))).
notation < "\infrule
(\infrule
non associative with precedence 19
for @{ 'lemma2_ko_2 $a $b $p ($H : $_) }.
interpretation "lemma2_ko_2" 'lemma2_ko_2 a b p H =
- (cast _ _ (show p (cast _ _
- (Lemma2 _ _ _ (cast _ _ H) (cast _ _ a) (cast _ _ b))))).
+ (cast ? ? (show p (cast ? ?
+ (Lemma2 ? ? ? (cast ? ? H) (cast ? ? a) (cast ? ? b))))).
notation < "\infrule
non associative with precedence 19
for @{ 'lemma2_ok_1 $a $b $p ($H : $_) }.
interpretation "lemma2_ok_1" 'lemma2_ok_1 a b p H =
- (show p (Lemma2 _ _ _ H a b)).
+ (show p (Lemma2 ? ? ? H a b)).
notation < "\infrule
(\infrule
non associative with precedence 19
for @{ 'lemma2_ok_2 $a $b $p ($H : $_) }.
interpretation "lemma2_ok_2" 'lemma2_ok_2 a b p H =
- (cast _ _ (show p (Lemma2 _ _ _ H a b))).
+ (cast ? ? (show p (Lemma2 ? ? ? H a b))).
notation > "'lem' 2 term 90 l term 90 p term 90 q" non associative with precedence 19
for @{ 'Lemma2 (hide_args ? $l : ?) (show $p ?) (show $q ?) }.
interpretation "lemma 2 KO" 'Lemma2 l p q =
- (cast _ _ (Lemma2 unit unit unit (cast (Imply unit (Imply unit unit)) _ l) (cast unit _ p) (cast unit _ q))).
-interpretation "lemma 2 OK" 'Lemma2 l p q = (Lemma2 _ _ _ l p q).
+ (cast ? ? (Lemma2 unit unit unit (cast (Imply unit (Imply unit unit)) ? l) (cast unit ? p) (cast unit ? q))).
+interpretation "lemma 2 OK" 'Lemma2 l p q = (Lemma2 ? ? ? l p q).
(* already proved lemma, 3 assumptions *)
definition Lemma3 : ΠA,B,C,D. (A ⇒ B ⇒ C ⇒ D) → A → B → C → D ≝
non associative with precedence 19
for @{ 'lemma3_ko_1 $a $b $c $p ($H : $_) }.
interpretation "lemma3_ko_1" 'lemma3_ko_1 a b c p H =
- (show p (cast _ _
- (Lemma3 _ _ _ _ (cast _ _ H) (cast _ _ a) (cast _ _ b) (cast _ _ c)))).
+ (show p (cast ? ?
+ (Lemma3 ? ? ? ? (cast ? ? H) (cast ? ? a) (cast ? ? b) (cast ? ? c)))).
notation < "\infrule
(\infrule
non associative with precedence 19
for @{ 'lemma3_ko_2 $a $b $c $p ($H : $_) }.
interpretation "lemma3_ko_2" 'lemma3_ko_2 a b c p H =
- (cast _ _ (show p (cast _ _
- (Lemma3 _ _ _ _ (cast _ _ H) (cast _ _ a) (cast _ _ b) (cast _ _ c))))).
+ (cast ? ? (show p (cast ? ?
+ (Lemma3 ? ? ? ? (cast ? ? H) (cast ? ? a) (cast ? ? b) (cast ? ? c))))).
notation < "\infrule
non associative with precedence 19
for @{ 'lemma3_ok_1 $a $b $c $p ($H : $_) }.
interpretation "lemma3_ok_1" 'lemma3_ok_1 a b c p H =
- (show p (Lemma3 _ _ _ _ H a b c)).
+ (show p (Lemma3 ? ? ? ? H a b c)).
notation < "\infrule
(\infrule
non associative with precedence 19
for @{ 'lemma3_ok_2 $a $b $c $p ($H : $_) }.
interpretation "lemma3_ok_2" 'lemma3_ok_2 a b c p H =
- (cast _ _ (show p (Lemma3 _ _ _ _ H a b c))).
+ (cast ? ? (show p (Lemma3 ? ? ? ? H a b c))).
notation > "'lem' 3 term 90 l term 90 p term 90 q term 90 r" non associative with precedence 19
for @{ 'Lemma3 (hide_args ? $l : ?) (show $p ?) (show $q ?) (show $r ?) }.
interpretation "lemma 3 KO" 'Lemma3 l p q r =
- (cast _ _ (Lemma3 unit unit unit unit (cast (Imply unit (Imply unit (Imply unit unit))) _ l) (cast unit _ p) (cast unit _ q) (cast unit _ r))).
-interpretation "lemma 3 OK" 'Lemma3 l p q r = (Lemma3 _ _ _ _ l p q r).
+ (cast ? ? (Lemma3 unit unit unit unit (cast (Imply unit (Imply unit (Imply unit unit))) ? l) (cast unit ? p) (cast unit ? q) (cast unit ? r))).
+interpretation "lemma 3 OK" 'Lemma3 l p q r = (Lemma3 ? ? ? ? l p q r).
\lambda A,B,C:Type.\lambda f:(B\to C).\lambda g:(A\to B).\lambda x:A.
f (g x).
-interpretation "function composition" 'compose f g = (compose _ _ _ f g).
+interpretation "function composition" 'compose f g = (compose ? ? ? f g).
definition injective: \forall A,B:Type.\forall f:A \to B.Prop
\def \lambda A,B.\lambda f.
right associative with precedence 47
for @{'append $l1 $l2 }.
-interpretation "nil" 'nil = (nil _).
-interpretation "cons" 'cons hd tl = (cons _ hd tl).
+interpretation "nil" 'nil = (nil ?).
+interpretation "cons" 'cons hd tl = (cons ? hd tl).
(* theorem test_notation: [O; S O; S (S O)] = O :: S O :: S (S O) :: []. *)
[ nil => []
| (cons hd tl) => tl].
-interpretation "append" 'append l1 l2 = (append _ l1 l2).
+interpretation "append" 'append l1 l2 = (append ? l1 l2).
theorem append_nil: \forall A:Type.\forall l:list A.l @ [] = l.
intros;
inductive ex (A:Type) (P:A \to Prop) : Prop \def
ex_intro: \forall x:A. P x \to ex A P.
-interpretation "exists" 'exists x = (ex _ x).
+interpretation "exists" 'exists x = (ex ? x).
inductive ex2 (A:Type) (P,Q:A \to Prop) : Prop \def
ex_intro2: \forall x:A. P x \to Q x \to ex2 A P Q.
inductive exT (A:Type) (P:A→CProp) : CProp ≝
ex_introT: ∀w:A. P w → exT A P.
-interpretation "CProp exists" 'exists x = (exT _ x).
+interpretation "CProp exists" 'exists x = (exT ? x).
notation "\ll term 19 a, break term 19 b \gg"
with precedence 90 for @{'dependent_pair $a $b}.
-interpretation "dependent pair" 'dependent_pair a b =
- (ex_introT _ _ a b).
+interpretation "dependent pair" 'dependent_pair a b = (ex_introT ?? a b).
definition pi1exT ≝ λA,P.λx:exT A P.match x with [ex_introT x _ ⇒ x].
-interpretation "exT \fst" 'pi1 = (pi1exT _ _).
-interpretation "exT \fst" 'pi1a x = (pi1exT _ _ x).
-interpretation "exT \fst" 'pi1b x y = (pi1exT _ _ x y).
+interpretation "exT \fst" 'pi1 = (pi1exT ? ?).
+interpretation "exT \fst" 'pi1a x = (pi1exT ? ? x).
+interpretation "exT \fst" 'pi1b x y = (pi1exT ? ? x y).
definition pi2exT ≝
λA,P.λx:exT A P.match x return λx.P (pi1exT ? ? x) with [ex_introT _ p ⇒ p].
-interpretation "exT \snd" 'pi2 = (pi2exT _ _).
-interpretation "exT \snd" 'pi2a x = (pi2exT _ _ x).
-interpretation "exT \snd" 'pi2b x y = (pi2exT _ _ x y).
+interpretation "exT \snd" 'pi2 = (pi2exT ? ?).
+interpretation "exT \snd" 'pi2a x = (pi2exT ? ? x).
+interpretation "exT \snd" 'pi2b x y = (pi2exT ? ? x y).
inductive exP (A:Type) (P:A→Prop) : CProp ≝
ex_introP: ∀w:A. P w → exP A P.
interpretation "dependent pair for Prop" 'dependent_pair a b =
- (ex_introP _ _ a b).
+ (ex_introP ?? a b).
-interpretation "CProp exists for Prop" 'exists x = (exP _ x).
+interpretation "CProp exists for Prop" 'exists x = (exP ? x).
definition pi1exP ≝ λA,P.λx:exP A P.match x with [ex_introP x _ ⇒ x].
-interpretation "exP \fst" 'pi1 = (pi1exP _ _).
-interpretation "exP \fst" 'pi1a x = (pi1exP _ _ x).
-interpretation "exP \fst" 'pi1b x y = (pi1exP _ _ x y).
+interpretation "exP \fst" 'pi1 = (pi1exP ? ?).
+interpretation "exP \fst" 'pi1a x = (pi1exP ? ? x).
+interpretation "exP \fst" 'pi1b x y = (pi1exP ? ? x y).
definition pi2exP ≝
λA,P.λx:exP A P.match x return λx.P (pi1exP ?? x) with [ex_introP _ p ⇒ p].
-interpretation "exP \snd" 'pi2 = (pi2exP _ _).
-interpretation "exP \snd" 'pi2a x = (pi2exP _ _ x).
-interpretation "exP \snd" 'pi2b x y = (pi2exP _ _ x y).
+interpretation "exP \snd" 'pi2 = (pi2exP ? ?).
+interpretation "exP \snd" 'pi2a x = (pi2exP ? ? x).
+interpretation "exP \snd" 'pi2b x y = (pi2exP ? ? x y).
inductive exT23 (A:Type) (P:A→CProp) (Q:A→CProp) (R:A→A→CProp) : CProp ≝
ex_introT23: ∀w,p:A. P w → Q p → R w p → exT23 A P Q R.
definition pi1exT23 ≝
λA,P,Q,R.λx:exT23 A P Q R.match x with [ex_introT23 x _ _ _ _ ⇒ x].
-interpretation "exT2 \fst" 'pi1 = (pi1exT23 _ _ _ _).
-interpretation "exT2 \fst" 'pi1a x = (pi1exT23 _ _ _ _ x).
-interpretation "exT2 \fst" 'pi1b x y = (pi1exT23 _ _ _ _ x y).
+interpretation "exT2 \fst" 'pi1 = (pi1exT23 ? ? ? ?).
+interpretation "exT2 \fst" 'pi1a x = (pi1exT23 ? ? ? ? x).
+interpretation "exT2 \fst" 'pi1b x y = (pi1exT23 ? ? ? ? x y).
definition pi2exT23 ≝
λA,P,Q,R.λx:exT23 A P Q R.match x with [ex_introT23 _ x _ _ _ ⇒ x].
-interpretation "exT2 \snd" 'pi2 = (pi2exT23 _ _ _ _).
-interpretation "exT2 \snd" 'pi2a x = (pi2exT23 _ _ _ _ x).
-interpretation "exT2 \snd" 'pi2b x y = (pi2exT23 _ _ _ _ x y).
+interpretation "exT2 \snd" 'pi2 = (pi2exT23 ? ? ? ?).
+interpretation "exT2 \snd" 'pi2a x = (pi2exT23 ? ? ? ? x).
+interpretation "exT2 \snd" 'pi2b x y = (pi2exT23 ? ? ? ? x y).
inductive exT2 (A:Type) (P,Q:A→CProp) : CProp ≝
ex_introT2: ∀w:A. P w → Q w → exT2 A P Q.
<property name="visible">True</property>
<property name="can_focus">True</property>
<property name="receives_default">True</property>
- <property name="label" translatable="yes">Implication (⇒_i)</property>
<property name="response_id">0</property>
+ <child>
+ <widget class="GtkLabel" id="label8">
+ <property name="visible">True</property>
+ <property name="label" translatable="yes">Implication (⇒<sub>i</sub>)</property>
+ <property name="use_markup">True</property>
+ </widget>
+ </child>
</widget>
</child>
<child>
<property name="visible">True</property>
<property name="can_focus">True</property>
<property name="receives_default">True</property>
- <property name="label" translatable="yes">Conjunction (∧_i)</property>
<property name="response_id">0</property>
+ <child>
+ <widget class="GtkLabel" id="label7">
+ <property name="visible">True</property>
+ <property name="label" translatable="yes">Conjunction (∧<sub>i</sub>)</property>
+ <property name="use_markup">True</property>
+ </widget>
+ </child>
</widget>
<packing>
<property name="position">1</property>
<property name="visible">True</property>
<property name="can_focus">True</property>
<property name="receives_default">True</property>
- <property name="label" translatable="yes">Disjunction left (∨_i_l)</property>
<property name="response_id">0</property>
+ <child>
+ <widget class="GtkLabel" id="label9">
+ <property name="visible">True</property>
+ <property name="label" translatable="yes">Disjunction left (∨<sub>i-l</sub>)</property>
+ <property name="use_markup">True</property>
+ </widget>
+ </child>
</widget>
<packing>
<property name="position">2</property>
<property name="visible">True</property>
<property name="can_focus">True</property>
<property name="receives_default">True</property>
- <property name="label" translatable="yes">Disjunction right (∨_i_r)</property>
<property name="response_id">0</property>
+ <child>
+ <widget class="GtkLabel" id="label10">
+ <property name="visible">True</property>
+ <property name="label" translatable="yes">Disjunction right (∨<sub>i-r</sub>)</property>
+ <property name="use_markup">True</property>
+ </widget>
+ </child>
</widget>
<packing>
<property name="position">3</property>
<property name="visible">True</property>
<property name="can_focus">True</property>
<property name="receives_default">True</property>
- <property name="label" translatable="yes">Negation (¬_i)</property>
<property name="response_id">0</property>
+ <child>
+ <widget class="GtkLabel" id="label11">
+ <property name="visible">True</property>
+ <property name="label" translatable="yes">Negation (¬<sub>i</sub>)</property>
+ <property name="use_markup">True</property>
+ </widget>
+ </child>
</widget>
<packing>
<property name="position">4</property>
<property name="visible">True</property>
<property name="can_focus">True</property>
<property name="receives_default">True</property>
- <property name="label" translatable="yes">Top (⊤_i)</property>
<property name="response_id">0</property>
+ <child>
+ <widget class="GtkLabel" id="label12">
+ <property name="visible">True</property>
+ <property name="label" translatable="yes">Top (⊤<sub>i</sub>)</property>
+ <property name="use_markup">True</property>
+ </widget>
+ </child>
</widget>
<packing>
<property name="position">5</property>
<property name="visible">True</property>
<property name="can_focus">True</property>
<property name="receives_default">True</property>
- <property name="label" translatable="yes">Universal (∀_i)</property>
<property name="response_id">0</property>
+ <child>
+ <widget class="GtkLabel" id="label20">
+ <property name="visible">True</property>
+ <property name="label" translatable="yes">Universal (∀<sub>i</sub>)</property>
+ <property name="use_markup">True</property>
+ </widget>
+ </child>
</widget>
<packing>
<property name="position">6</property>
<property name="visible">True</property>
<property name="can_focus">True</property>
<property name="receives_default">True</property>
- <property name="label" translatable="yes">Existential (∃_i)</property>
<property name="response_id">0</property>
+ <child>
+ <widget class="GtkLabel" id="label21">
+ <property name="visible">True</property>
+ <property name="label" translatable="yes">Existential (∃<sub>i</sub>)</property>
+ <property name="use_markup">True</property>
+ </widget>
+ </child>
</widget>
<packing>
<property name="position">7</property>
<property name="visible">True</property>
<property name="can_focus">True</property>
<property name="receives_default">True</property>
- <property name="label" translatable="yes">Implication (⇒_e)</property>
<property name="response_id">0</property>
+ <child>
+ <widget class="GtkLabel" id="label22">
+ <property name="visible">True</property>
+ <property name="label" translatable="yes">Implication (⇒<sub>e</sub>)</property>
+ <property name="use_markup">True</property>
+ </widget>
+ </child>
</widget>
</child>
<child>
<property name="visible">True</property>
<property name="can_focus">True</property>
<property name="receives_default">True</property>
- <property name="label" translatable="yes">Conjunction left (∧_e_l)</property>
<property name="response_id">0</property>
+ <child>
+ <widget class="GtkLabel" id="label23">
+ <property name="visible">True</property>
+ <property name="label" translatable="yes">Conjunction left (∧<sub>e-l</sub>)</property>
+ <property name="use_markup">True</property>
+ </widget>
+ </child>
</widget>
<packing>
<property name="position">1</property>
<property name="visible">True</property>
<property name="can_focus">True</property>
<property name="receives_default">True</property>
- <property name="label" translatable="yes">Conjunction right (∧_e_r)</property>
<property name="response_id">0</property>
+ <child>
+ <widget class="GtkLabel" id="label24">
+ <property name="visible">True</property>
+ <property name="label" translatable="yes">Conjunction right (∧<sub>e-r</sub>)</property>
+ <property name="use_markup">True</property>
+ </widget>
+ </child>
</widget>
<packing>
<property name="position">2</property>
<property name="visible">True</property>
<property name="can_focus">True</property>
<property name="receives_default">True</property>
- <property name="label" translatable="yes">Disjunction (∨_e)</property>
<property name="response_id">0</property>
+ <child>
+ <widget class="GtkLabel" id="label27">
+ <property name="visible">True</property>
+ <property name="label" translatable="yes">Disjunction (∨<sub>e</sub>)</property>
+ <property name="use_markup">True</property>
+ </widget>
+ </child>
</widget>
<packing>
<property name="position">3</property>
<property name="visible">True</property>
<property name="can_focus">True</property>
<property name="receives_default">True</property>
- <property name="label" translatable="yes">Negation (¬_e)</property>
<property name="response_id">0</property>
+ <child>
+ <widget class="GtkLabel" id="label31">
+ <property name="visible">True</property>
+ <property name="label" translatable="yes">Negation (¬<sub>e</sub>)</property>
+ <property name="use_markup">True</property>
+ </widget>
+ </child>
</widget>
<packing>
<property name="position">4</property>
<property name="visible">True</property>
<property name="can_focus">True</property>
<property name="receives_default">True</property>
- <property name="label" translatable="yes">Bottom (⊥_e)</property>
<property name="response_id">0</property>
+ <child>
+ <widget class="GtkLabel" id="label33">
+ <property name="visible">True</property>
+ <property name="label" translatable="yes">Bottom (⊥<sub>e</sub>)</property>
+ <property name="use_markup">True</property>
+ </widget>
+ </child>
</widget>
<packing>
<property name="position">5</property>
<property name="visible">True</property>
<property name="can_focus">True</property>
<property name="receives_default">True</property>
- <property name="label" translatable="yes">Universal (∀_e)</property>
<property name="response_id">0</property>
+ <child>
+ <widget class="GtkLabel" id="label34">
+ <property name="visible">True</property>
+ <property name="label" translatable="yes">Universal (∀<sub>e</sub>)</property>
+ <property name="use_markup">True</property>
+ </widget>
+ </child>
</widget>
<packing>
<property name="position">6</property>
<property name="visible">True</property>
<property name="can_focus">True</property>
<property name="receives_default">True</property>
- <property name="label" translatable="yes">Existential (∃_e)</property>
<property name="response_id">0</property>
+ <child>
+ <widget class="GtkLabel" id="label35">
+ <property name="visible">True</property>
+ <property name="label" translatable="yes">Existential (∃<sub>e</sub>)</property>
+ <property name="use_markup">True</property>
+ </widget>
+ </child>
</widget>
<packing>
<property name="position">7</property>
else main#tacticsButtonsHandlebox#misc#hide ())
~check:main#menuitemPalette;
connect_button main#butImpl_intro
- (fun () -> source_buffer#insert "apply rule (⇒_i […] (…));\n");
+ (fun () -> source_buffer#insert "apply rule (⇒#i […] (…));\n");
connect_button main#butAnd_intro
(fun () -> source_buffer#insert
- "apply rule (∧_i (…) (…));\n\t[\n\t|\n\t]\n");
+ "apply rule (∧#i (…) (…));\n\t[\n\t|\n\t]\n");
connect_button main#butOr_intro_left
- (fun () -> source_buffer#insert "apply rule (∨_i_l (…));\n");
+ (fun () -> source_buffer#insert "apply rule (∨#i_l (…));\n");
connect_button main#butOr_intro_right
- (fun () -> source_buffer#insert "apply rule (∨_i_r (…));\n");
+ (fun () -> source_buffer#insert "apply rule (∨#i_r (…));\n");
connect_button main#butNot_intro
- (fun () -> source_buffer#insert "apply rule (¬_i […] (…));\n");
+ (fun () -> source_buffer#insert "apply rule (¬#i […] (…));\n");
connect_button main#butTop_intro
- (fun () -> source_buffer#insert "apply rule (⊤_i);\n");
+ (fun () -> source_buffer#insert "apply rule (⊤#i);\n");
connect_button main#butImpl_elim
(fun () -> source_buffer#insert
- "apply rule (⇒_e (…) (…));\n\t[\n\t|\n\t]\n");
+ "apply rule (⇒#e (…) (…));\n\t[\n\t|\n\t]\n");
connect_button main#butAnd_elim_left
- (fun () -> source_buffer#insert "apply rule (∧_e_l (…));\n");
+ (fun () -> source_buffer#insert "apply rule (∧#e_l (…));\n");
connect_button main#butAnd_elim_right
- (fun () -> source_buffer#insert "apply rule (∧_e_r (…));\n");
+ (fun () -> source_buffer#insert "apply rule (∧#e_r (…));\n");
connect_button main#butOr_elim
(fun () -> source_buffer#insert
- "apply rule (∨_e (…) […] (…) […] (…));\n\t[\n\t|\n\t|\n\t]\n");
+ "apply rule (∨#e (…) […] (…) […] (…));\n\t[\n\t|\n\t|\n\t]\n");
connect_button main#butNot_elim
(fun () -> source_buffer#insert
- "apply rule (¬_e (…) (…));\n\t[\n\t|\n\t]\n");
+ "apply rule (¬#e (…) (…));\n\t[\n\t|\n\t]\n");
connect_button main#butBot_elim
- (fun () -> source_buffer#insert "apply rule (⊥_e (…));\n");
+ (fun () -> source_buffer#insert "apply rule (⊥#e (…));\n");
connect_button main#butRAA
(fun () -> source_buffer#insert "apply rule (RAA […] (…));\n");
connect_button main#butUseLemma
(fun () -> source_buffer#insert "apply rule (discharge […]);\n");
connect_button main#butForall_intro
- (fun () -> source_buffer#insert "apply rule (∀_i {…} (…));\n");
+ (fun () -> source_buffer#insert "apply rule (∀#i {…} (…));\n");
connect_button main#butForall_elim
- (fun () -> source_buffer#insert "apply rule (∀_e {…} (…));\n");
+ (fun () -> source_buffer#insert "apply rule (∀#e {…} (…));\n");
connect_button main#butExists_intro
- (fun () -> source_buffer#insert "apply rule (∃_i {…} (…));\n");
+ (fun () -> source_buffer#insert "apply rule (∃#i {…} (…));\n");
connect_button main#butExists_elim
(fun () -> source_buffer#insert
- "apply rule (∃_e (…) {…} […] (…));\n\t[\n\t|\n\t]\n");
+ "apply rule (∃#e (…) {…} […] (…));\n\t[\n\t|\n\t]\n");
(* TO BE REMOVED *)
TPTP/Veloci/GRP595-1.p.ma logic/equality.ma
TPTP/Veloci/GRP570-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/PLA011-2.ma logic/equality.ma
dependent_guarded_bove_capretta.ma nat/nat.ma
-TPTP/Veloci/SYN083-1.p.ma logic/equality.ma
interactive/test5.ma
+TPTP/Veloci/SYN083-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/SWV014-1.ma logic/equality.ma
TPTP/Veloci/COL008-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL063-1.ma logic/equality.ma
+TPTP/Unsatisfiable/NUM017-1.ma logic/equality.ma
TPTP/Veloci/GRP604-1.p.ma logic/equality.ma
record.ma
TPTP/Veloci/GRP516-1.p.ma logic/equality.ma
TPTP/Veloci/GRP612-1.p.ma logic/equality.ma
TPTP/Veloci/COL024-1.p.ma logic/equality.ma
TPTP/Veloci/GRP139-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL167-1.ma logic/equality.ma
TPTP/Veloci/ROB010-1.p.ma logic/equality.ma
TPTP/Veloci/BOO006-2.p.ma logic/equality.ma
TPTP/Veloci/LCL157-1.p.ma logic/equality.ma
TPTP/Veloci/GRP155-1.p.ma logic/equality.ma
TPTP/Veloci/BOO004-4.p.ma logic/equality.ma
TPTP/Veloci/LCL140-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP598-1.p.ma logic/equality.ma
TPTP/Veloci/GRP573-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP485-1.p.ma logic/equality.ma
+TPTP/Veloci/GRP598-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL393-1.ma logic/equality.ma
TPTP/Veloci/GRP460-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP188-1.p.ma logic/equality.ma
TPTP/Veloci/GRP163-1.p.ma logic/equality.ma
+TPTP/Veloci/GRP485-1.p.ma logic/equality.ma
+TPTP/Veloci/GRP188-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL253-1.ma logic/equality.ma
TPTP/Veloci/BOO012-4.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL375-1.ma logic/equality.ma
TPTP/Veloci/GRP581-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP493-1.p.ma logic/equality.ma
TPTP/Veloci/COL063-3.p.ma logic/equality.ma
TPTP/Veloci/LCL112-2.p.ma logic/equality.ma
+TPTP/Veloci/GRP493-1.p.ma logic/equality.ma
formal_topology.ma logic/connectives.ma
TPTP/Veloci/RNG023-7.p.ma logic/equality.ma
replace.ma coq.ma
test2.ma coq.ma
-TPTP/Veloci/COL064-7.p.ma logic/equality.ma
TPTP/Veloci/BOO009-2.p.ma logic/equality.ma
+TPTP/Veloci/COL064-7.p.ma logic/equality.ma
TPTP/Veloci/GRP510-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL037-1.ma logic/equality.ma
TPTP/Veloci/COL010-1.p.ma logic/equality.ma
-TPTP/Veloci/BOO017-2.p.ma logic/equality.ma
-TPTP/Veloci/LCL135-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/SYN615-1.ma logic/equality.ma
TPTP/Veloci/COL061-2.p.ma logic/equality.ma
+TPTP/Veloci/LCL135-1.p.ma logic/equality.ma
+TPTP/Veloci/BOO017-2.p.ma logic/equality.ma
+TPTP/Unsatisfiable/SYN600-1.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL019-1.ma logic/equality.ma
elim_pattern.ma nat/plus.ma
assumption.ma coq.ma
-interactive/grafite.ma
TPTP/Veloci/GRP551-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP463-1.p.ma logic/equality.ma
+interactive/grafite.ma
TPTP/Veloci/GRP141-1.p.ma logic/equality.ma
+TPTP/Veloci/GRP463-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL001-1.ma logic/equality.ma
fix_betareduction.ma coq.ma
+TPTP/Unsatisfiable/SYN704-1.ma logic/equality.ma
TPTP/Veloci/GRP606-1.p.ma logic/equality.ma
coercions_dependent.ma decidable_kit/list_aux.ma list/list.ma nat/nat.ma
-TPTP/Veloci/GRP584-1.p.ma logic/equality.ma
TPTP/Veloci/GRP518-1.p.ma logic/equality.ma
-paramodulation/boolean_algebra.ma coq.ma
-TPTP/Veloci/GRP496-1.p.ma logic/equality.ma
+TPTP/Veloci/GRP584-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/PLA016-1.ma logic/equality.ma
TPTP/Veloci/COL084-1.p.ma logic/equality.ma
TPTP/Veloci/GRP174-1.p.ma logic/equality.ma
+paramodulation/boolean_algebra.ma coq.ma
+TPTP/Veloci/GRP496-1.p.ma logic/equality.ma
TPTP/Veloci/RNG011-5.p.ma logic/equality.ma
TPTP/Veloci/GRP614-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/PLA001-1.ma logic/equality.ma
inversion.ma coq.ma
TPTP/Veloci/GRP592-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP182-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL105-1.ma logic/equality.ma
bad_tests/baseuri.ma
+TPTP/Veloci/GRP182-1.p.ma logic/equality.ma
TPTP/Veloci/GRP116-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL230-1.ma logic/equality.ma
coercions.ma nat/compare.ma nat/times.ma
+TPTP/Unsatisfiable/LCL227-1.ma logic/equality.ma
TPTP/Veloci/GRP011-4.p.ma logic/equality.ma
TPTP/Veloci/GRP149-1.p.ma logic/equality.ma
TPTP/Veloci/COL013-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP567-1.p.ma logic/equality.ma
TPTP/Veloci/COL064-2.p.ma logic/equality.ma
+TPTP/Veloci/GRP567-1.p.ma logic/equality.ma
comments.ma coq.ma
-TPTP/Veloci/COL021-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL032-1.ma logic/equality.ma
TPTP/Veloci/BOO006-4.p.ma logic/equality.ma
+TPTP/Veloci/COL021-1.p.ma logic/equality.ma
TPTP/Veloci/BOO003-2.p.ma logic/equality.ma
TPTP/Veloci/RNG024-6.p.ma logic/equality.ma
bool.ma coq.ma
TPTP/Veloci/BOO011-2.p.ma logic/equality.ma
TPTP/Veloci/LCL154-1.p.ma logic/equality.ma
TPTP/Veloci/LCL114-2.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL423-1.ma logic/equality.ma
letrec.ma
-TPTP/Veloci/BOO001-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL395-1.ma logic/equality.ma
+TPTP/Unsatisfiable/PLA011-1.ma logic/equality.ma
TPTP/Veloci/GRP160-1.p.ma logic/equality.ma
+TPTP/Veloci/BOO001-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/PLA008-1.ma logic/equality.ma
TPTP/Veloci/GRP600-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL377-1.ma logic/equality.ma
+TPTP/Unsatisfiable/PLA005-2.ma logic/equality.ma
TPTP/Veloci/GRP512-1.p.ma logic/equality.ma
TPTP/Veloci/COL063-5.p.ma logic/equality.ma
continuationals.ma coq.ma
TPTP/Veloci/GRP490-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/SYN653-1.ma logic/equality.ma
TPTP/Veloci/BOO034-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/ANA003-2.ma logic/equality.ma
+ng_pts.ma
fold.ma coq.ma
-TPTP/Veloci/GRP545-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL084-3.ma logic/equality.ma
TPTP/Veloci/GRP520-1.p.ma logic/equality.ma
+TPTP/Veloci/GRP545-1.p.ma logic/equality.ma
+a.ma nat/nat.ma
TPTP/Veloci/GRP457-1.p.ma logic/equality.ma
-TPTP/Veloci/BOO009-4.p.ma logic/equality.ma
-TPTP/Veloci/COL064-9.p.ma logic/equality.ma
TPTP/Veloci/LAT040-1.p.ma logic/equality.ma
+TPTP/Veloci/COL064-9.p.ma logic/equality.ma
+TPTP/Veloci/BOO009-4.p.ma logic/equality.ma
TPTP/Veloci/GRP578-1.p.ma logic/equality.ma
change.ma coq.ma
TPTP/Veloci/BOO075-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP168-1.p.ma logic/equality.ma
TPTP/Veloci/GRP143-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP586-1.p.ma logic/equality.ma
+TPTP/Veloci/GRP168-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/SYN617-1.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL024-1.ma logic/equality.ma
TPTP/Veloci/GRP561-1.p.ma logic/equality.ma
+TPTP/Veloci/GRP586-1.p.ma logic/equality.ma
TPTP/Veloci/GRP498-1.p.ma logic/equality.ma
TPTP/Veloci/GRP176-1.p.ma logic/equality.ma
overred.ma logic/equality.ma
TPTP/Veloci/COL007-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/PLA021-1.ma logic/equality.ma
+TPTP/Unsatisfiable/SYN556-1.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL125-1.ma logic/equality.ma
+TPTP/Unsatisfiable/PLA018-1.ma logic/equality.ma
TPTP/Veloci/GRP603-1.p.ma logic/equality.ma
TPTP/Veloci/COL058-2.p.ma logic/equality.ma
TPTP/Veloci/GRP515-1.p.ma logic/equality.ma
-TPTP/Veloci/COL015-1.p.ma logic/equality.ma
TPTP/Veloci/RNG007-4.p.ma logic/equality.ma
+TPTP/Veloci/COL015-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL085-1.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL369-1.ma logic/equality.ma
TPTP/Veloci/GRP548-1.p.ma logic/equality.ma
TPTP/Veloci/GRP182-3.p.ma logic/equality.ma
TPTP/Veloci/COL048-1.p.ma logic/equality.ma
TPTP/Veloci/BOO005-2.p.ma logic/equality.ma
TPTP/Veloci/GRP556-1.p.ma logic/equality.ma
TPTP/Veloci/GRP146-1.p.ma logic/equality.ma
-TPTP/Veloci/COL064-4.p.ma logic/equality.ma
TPTP/Veloci/LCL156-1.p.ma logic/equality.ma
+TPTP/Veloci/COL064-4.p.ma logic/equality.ma
contradiction.ma coq.ma
TPTP/Veloci/GRP564-1.p.ma logic/equality.ma
inversion2.ma coq.ma
TPTP/Veloci/GRP154-1.p.ma logic/equality.ma
-TPTP/Veloci/ROB009-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/SYN599-1.ma logic/equality.ma
TPTP/Veloci/LCL164-1.p.ma logic/equality.ma
+TPTP/Veloci/ROB009-1.p.ma logic/equality.ma
TPTP/Veloci/BOO003-4.p.ma logic/equality.ma
ng_elim.ma
TPTP/Veloci/GRP597-1.p.ma logic/equality.ma
TPTP/Veloci/GRP572-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP162-1.p.ma logic/equality.ma
TPTP/Veloci/GRP484-1.p.ma logic/equality.ma
+TPTP/Veloci/GRP162-1.p.ma logic/equality.ma
fix_che_non_passa_ma_dovrebbe.ma list/list.ma nat/nat.ma
TPTP/Veloci/BOO011-4.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL425-1.ma logic/equality.ma
TPTP/Veloci/GRP580-1.p.ma logic/equality.ma
TPTP/Veloci/COL062-3.p.ma logic/equality.ma
TPTP/Veloci/GRP492-1.p.ma logic/equality.ma
cut.ma coq.ma
-TPTP/Veloci/COL018-1.p.ma logic/equality.ma
TPTP/Veloci/LAT034-1.p.ma logic/equality.ma
-tinycals.ma logic/connectives.ma
+TPTP/Veloci/COL018-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/PLA013-1.ma logic/equality.ma
coercions_contravariant.ma logic/equality.ma nat/nat.ma
+tinycals.ma logic/connectives.ma
interactive/automatic_insertion.ma
clearbody.ma coq.ma
+TPTP/Unsatisfiable/SYN655-1.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL062-1.ma logic/equality.ma
+TPTP/Veloci/BOO016-2.p.ma logic/equality.ma
TPTP/Veloci/LCL134-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/SYN640-1.ma logic/equality.ma
TPTP/Veloci/COL060-2.p.ma logic/equality.ma
-TPTP/Veloci/BOO016-2.p.ma logic/equality.ma
-TPTP/Veloci/LAT039-2.p.ma logic/equality.ma
TPTP/Veloci/GRP542-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP186-4.p.ma logic/equality.ma
+TPTP/Veloci/LAT039-2.p.ma logic/equality.ma
TPTP/Veloci/GRP454-1.p.ma logic/equality.ma
+TPTP/Veloci/GRP186-4.p.ma logic/equality.ma
TPTP/Veloci/GRP157-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP509-1.p.ma logic/equality.ma
TPTP/Veloci/GRP550-1.p.ma logic/equality.ma
-ng_commands.ma nat/plus.ma
+TPTP/Veloci/GRP509-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL166-1.ma logic/equality.ma
+ng_commands.ma ng_pts.ma
TPTP/Veloci/GRP487-1.p.ma logic/equality.ma
-paramodulation/BOO075-1.ma
TPTP/Veloci/GRP605-1.p.ma logic/equality.ma
+paramodulation/BOO075-1.ma
fguidi.ma logic/connectives.ma nat/nat.ma
TPTP/Veloci/GRP583-1.p.ma logic/equality.ma
TPTP/Veloci/GRP517-1.p.ma logic/equality.ma
+TPTP/Veloci/GRP173-1.p.ma logic/equality.ma
TPTP/Veloci/GRP495-1.p.ma logic/equality.ma
TPTP/Veloci/COL083-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP173-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/SYN711-1.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL005-1.ma logic/equality.ma
+TPTP/Unsatisfiable/SYN708-1.ma logic/equality.ma
+TPTP/Unsatisfiable/PLA023-1.ma logic/equality.ma
coercions_propagation.ma logic/connectives.ma nat/orders.ma
TPTP/Veloci/LAT045-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP010-4.p.ma logic/equality.ma
TPTP/Veloci/GRP558-1.p.ma logic/equality.ma
+TPTP/Veloci/GRP010-4.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL249-1.ma logic/equality.ma
+TPTP/Unsatisfiable/PLA005-1.ma logic/equality.ma
interactive/test_instance.ma
TPTP/Veloci/COL012-1.p.ma logic/equality.ma
TPTP/Veloci/COL063-2.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL084-2.ma logic/equality.ma
first.ma
TPTP/Veloci/COL045-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/SYN647-1.ma logic/equality.ma
TPTP/Veloci/BOO005-4.p.ma logic/equality.ma
interactive/test7.ma
bad_induction.ma logic/equality.ma nat/nat.ma
TPTP/Veloci/GRP608-1.p.ma logic/equality.ma
TPTP/Veloci/BOO010-2.p.ma logic/equality.ma
TPTP/Veloci/LCL153-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/SYN614-1.ma logic/equality.ma
pullback.ma logic/equality.ma
TPTP/Veloci/COL086-1.p.ma logic/equality.ma
TPTP/Veloci/GRP151-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL021-1.ma logic/equality.ma
TPTP/Veloci/LCL161-1.p.ma logic/equality.ma
TPTP/Veloci/GRP616-1.p.ma logic/equality.ma
TPTP/Veloci/GRP023-2.p.ma logic/equality.ma
generalize.ma coq.ma
TPTP/Veloci/GRP511-1.p.ma logic/equality.ma
TPTP/Veloci/GRP192-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL122-1.ma logic/equality.ma
+TPTP/Unsatisfiable/PLA015-1.ma logic/equality.ma
TPTP/Veloci/GRP569-1.p.ma logic/equality.ma
TPTP/Veloci/GRP544-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL119-1.ma logic/equality.ma
dependent_type_inference.ma nat/nat.ma
TPTP/Veloci/GRP456-1.p.ma logic/equality.ma
TPTP/Veloci/GRP159-1.p.ma logic/equality.ma
TPTP/Veloci/ROB030-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/PLA009-2.ma logic/equality.ma
TPTP/Veloci/GRP577-1.p.ma logic/equality.ma
TPTP/Veloci/GRP552-1.p.ma logic/equality.ma
TPTP/Veloci/BOO013-2.p.ma logic/equality.ma
+TPTP/Unsatisfiable/SYN639-1.ma logic/equality.ma
rewrite.ma coq.ma
-TPTP/Veloci/GRP497-1.p.ma logic/equality.ma
bad_tests/auto.ma coq.ma
+TPTP/Veloci/GRP497-1.p.ma logic/equality.ma
diabolic_fix.ma nat/nat.ma
TPTP/Veloci/GRP001-2.p.ma logic/equality.ma
test4.ma coq.ma
TPTP/Veloci/GRP514-1.p.ma logic/equality.ma
TPTP/Veloci/LCL139-1.p.ma logic/equality.ma
TPTP/Veloci/GRP547-1.p.ma logic/equality.ma
+TPTP/Veloci/COL004-3.p.ma logic/equality.ma
TPTP/Veloci/GRP459-1.p.ma logic/equality.ma
TPTP/Veloci/BOO069-1.p.ma logic/equality.ma
TPTP/Veloci/GRP137-1.p.ma logic/equality.ma
-TPTP/Veloci/COL004-3.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL394-1.ma logic/equality.ma
+TPTP/Unsatisfiable/PLA010-1.ma logic/equality.ma
TPTP/Veloci/GRP188-2.p.ma logic/equality.ma
apply.ma coq.ma
+TPTP/Unsatisfiable/PLA007-1.ma logic/equality.ma
TPTP/Veloci/GRP467-1.p.ma logic/equality.ma
TPTP/Veloci/GRP145-1.p.ma logic/equality.ma
-TPTP/Veloci/COL063-4.p.ma logic/equality.ma
TPTP/Veloci/LCL155-1.p.ma logic/equality.ma
+TPTP/Veloci/COL063-4.p.ma logic/equality.ma
+TPTP/Unsatisfiable/PLA004-2.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL074-1.ma logic/equality.ma
TPTP/Veloci/GRP588-1.p.ma logic/equality.ma
TPTP/Veloci/GRP153-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/SYN649-1.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL218-1.ma logic/equality.ma
paramodulation.ma coq.ma
TPTP/Veloci/GRP596-1.p.ma logic/equality.ma
coercions_nonuniform.ma
TPTP/Veloci/LDA001-1.p.ma logic/equality.ma
TPTP/Veloci/COL050-1.p.ma logic/equality.ma
TPTP/Veloci/BOO010-4.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL038-1.ma logic/equality.ma
TPTP/Veloci/LCL110-2.p.ma logic/equality.ma
TPTP/Veloci/GRP491-1.p.ma logic/equality.ma
TPTP/Veloci/COL061-3.p.ma logic/equality.ma
TPTP/Veloci/COL017-1.p.ma logic/equality.ma
TPTP/Veloci/GRP613-1.p.ma logic/equality.ma
compose.ma logic/equality.ma
-TPTP/Veloci/COL025-1.p.ma logic/equality.ma
TPTP/Veloci/GRP115-1.p.ma logic/equality.ma
-hard_refine.ma coq.ma
-letrecand.ma nat/nat.ma
+TPTP/Veloci/COL025-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL002-1.ma logic/equality.ma
unifhint.ma list/list.ma nat/nat.ma nat/plus.ma
+letrecand.ma nat/nat.ma
+hard_refine.ma coq.ma
paramodulation/group.ma coq.ma
+TPTP/Unsatisfiable/LCL124-1.ma logic/equality.ma
TPTP/Veloci/LCL158-1.p.ma logic/equality.ma
TPTP/Veloci/LCL133-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP566-1.p.ma logic/equality.ma
TPTP/Veloci/GRP541-1.p.ma logic/equality.ma
+TPTP/Veloci/GRP566-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/PLA014-2.ma logic/equality.ma
decl.ma nat/orders.ma nat/times.ma
TPTP/Veloci/GRP156-1.p.ma logic/equality.ma
mysql_escaping.ma
TPTP/Veloci/LCL141-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL231-1.ma logic/equality.ma
+TPTP/Unsatisfiable/PUZ040-1.ma logic/equality.ma
TPTP/Veloci/GRP182-2.p.ma logic/equality.ma
-TPTP/Veloci/LDA007-3.p.ma logic/equality.ma
+TPTP/Veloci/GRP574-1.p.ma logic/equality.ma
+TPTP/Veloci/GRP599-1.p.ma logic/equality.ma
interactive/drop.ma
interactive/test6.ma
-TPTP/Veloci/GRP599-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP574-1.p.ma logic/equality.ma
+TPTP/Veloci/LDA007-3.p.ma logic/equality.ma
unfold.ma coq.ma
TPTP/Veloci/GRP486-1.p.ma logic/equality.ma
TPTP/Veloci/GRP189-1.p.ma logic/equality.ma
TPTP/Veloci/BOO071-1.p.ma logic/equality.ma
TPTP/Veloci/BOO013-4.p.ma logic/equality.ma
TPTP/Veloci/GRP582-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP494-1.p.ma logic/equality.ma
TPTP/Veloci/COL064-3.p.ma logic/equality.ma
TPTP/Veloci/LCL113-2.p.ma logic/equality.ma
+TPTP/Veloci/GRP494-1.p.ma logic/equality.ma
TPTP/Veloci/RNG008-4.p.ma logic/equality.ma
demodulation_matita.ma nat/minus.ma
TPTP/Veloci/ROB002-1.p.ma logic/equality.ma
tacticals.ma
-TPTP/Veloci/GRP001-4.p.ma logic/equality.ma
TPTP/Veloci/GRP590-1.p.ma logic/equality.ma
TPTP/Veloci/GRP549-1.p.ma logic/equality.ma
+TPTP/Veloci/GRP001-4.p.ma logic/equality.ma
+TPTP/Unsatisfiable/SYN598-1.ma logic/equality.ma
applys.ma nat/div_and_mod.ma nat/factorial.ma nat/primes.ma
TPTP/Veloci/RNG024-7.p.ma logic/equality.ma
+TPTP/Unsatisfiable/RNG004-3.ma logic/equality.ma
injection.ma coq.ma
second.ma first.ma
TPTP/Veloci/COL062-2.p.ma logic/equality.ma
match_inference.ma
-simpl.ma coq.ma
+TPTP/Unsatisfiable/PLA012-1.ma logic/equality.ma
unifhint_simple.ma logic/equality.ma
+simpl.ma coq.ma
+TPTP/Unsatisfiable/PLA009-1.ma logic/equality.ma
+TPTP/Unsatisfiable/PUZ050-1.ma logic/equality.ma
TPTP/Veloci/GRP142-1.p.ma logic/equality.ma
TPTP/Veloci/COL063-6.p.ma logic/equality.ma
+TPTP/Unsatisfiable/SYN654-1.ma logic/equality.ma
third.ma first.ma second.ma
+TPTP/Unsatisfiable/LCL061-1.ma logic/equality.ma
TPTP/Veloci/GRP560-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP150-1.p.ma logic/equality.ma
TPTP/Veloci/COL085-1.p.ma logic/equality.ma
+TPTP/Veloci/GRP150-1.p.ma logic/equality.ma
TPTP/Veloci/GRP615-1.p.ma logic/equality.ma
TPTP/Veloci/LAT039-1.p.ma logic/equality.ma
luo.ma Z/times.ma logic/equality.ma
naiveparamod.ma logic/equality.ma
TPTP/Veloci/GRP176-2.p.ma logic/equality.ma
TPTP/Veloci/COL014-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP568-1.p.ma logic/equality.ma
TPTP/Veloci/GRP543-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP455-1.p.ma logic/equality.ma
+TPTP/Veloci/GRP568-1.p.ma logic/equality.ma
TPTP/Veloci/GRP158-1.p.ma logic/equality.ma
+TPTP/Veloci/GRP455-1.p.ma logic/equality.ma
TPTP/Veloci/COL022-1.p.ma logic/equality.ma
TPTP/Veloci/GRP576-1.p.ma logic/equality.ma
TPTP/Veloci/BOO004-2.p.ma logic/equality.ma
+TPTP/Unsatisfiable/SYN707-1.ma logic/equality.ma
TPTP/Veloci/GRP488-1.p.ma logic/equality.ma
TPTP/Veloci/COL058-3.p.ma logic/equality.ma
+TPTP/Unsatisfiable/PLA019-1.ma logic/equality.ma
dependent_injection.ma coq.ma
TPTP/Veloci/BOO012-2.p.ma logic/equality.ma
+TPTP/Unsatisfiable/PLA004-1.ma logic/equality.ma
TPTP/Veloci/LCL115-2.p.ma logic/equality.ma
absurd.ma coq.ma
+TPTP/Unsatisfiable/PUZ042-1.ma logic/equality.ma
TPTP/Veloci/GRP182-4.p.ma logic/equality.ma
+TPTP/Unsatisfiable/PUZ039-1.ma logic/equality.ma
+TPTP/Unsatisfiable/SYN646-1.ma logic/equality.ma
destruct.ma datatypes/constructors.ma logic/equality.ma nat/nat.ma
-TPTP/Veloci/COL064-5.p.ma logic/equality.ma
+TPTP/Unsatisfiable/SYN631-1.ma logic/equality.ma
TPTP/Veloci/GRP513-1.p.ma logic/equality.ma
+TPTP/Veloci/COL064-5.p.ma logic/equality.ma
+TPTP/Unsatisfiable/SYN628-1.ma logic/equality.ma
TPTP/Veloci/LAT033-1.p.ma logic/equality.ma
TPTP/Veloci/GRP546-1.p.ma logic/equality.ma
-TPTP/Veloci/GRP136-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/LCL020-1.ma logic/equality.ma
TPTP/Veloci/GRP458-1.p.ma logic/equality.ma
+TPTP/Veloci/GRP136-1.p.ma logic/equality.ma
coercions_open.ma logic/equality.ma nat/nat.ma
paramodulation/irratsqrt2.ma nat/minus.ma nat/times.ma
TPTP/Veloci/GRP144-1.p.ma logic/equality.ma
TPTP/Veloci/LAT008-1.p.ma logic/equality.ma
demodulation_coq.ma coq.ma
TPTP/Veloci/GRP562-1.p.ma logic/equality.ma
+TPTP/Unsatisfiable/PLA014-1.ma logic/equality.ma
TPTP/Veloci/GRP152-1.p.ma logic/equality.ma
Z/times.ma
coq.ma